--- a/src/HOL/Algebra/AbelCoset.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Algebra/AbelCoset.thy Thu May 26 17:51:22 2016 +0200
@@ -10,7 +10,7 @@
subsubsection \<open>Definitions\<close>
-text \<open>Hiding @{text "<+>"} from @{theory Sum_Type} until I come
+text \<open>Hiding \<open><+>\<close> from @{theory Sum_Type} until I come
up with better syntax here\<close>
no_notation Sum_Type.Plus (infixr "<+>" 65)
@@ -41,12 +41,12 @@
definition
A_FactGroup :: "[('a,'b) ring_scheme, 'a set] \<Rightarrow> ('a set) monoid" (infixl "A'_Mod" 65)
- --\<open>Actually defined for groups rather than monoids\<close>
+ \<comment>\<open>Actually defined for groups rather than monoids\<close>
where "A_FactGroup G H = FactGroup \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr> H"
definition
a_kernel :: "('a, 'm) ring_scheme \<Rightarrow> ('b, 'n) ring_scheme \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a set"
- --\<open>the kernel of a homomorphism (additive)\<close>
+ \<comment>\<open>the kernel of a homomorphism (additive)\<close>
where "a_kernel G H h =
kernel \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>
\<lparr>carrier = carrier H, mult = add H, one = zero H\<rparr> h"
@@ -218,7 +218,7 @@
subsubsection \<open>Additive subgroups are normal\<close>
-text \<open>Every subgroup of an @{text "abelian_group"} is normal\<close>
+text \<open>Every subgroup of an \<open>abelian_group\<close> is normal\<close>
locale abelian_subgroup = additive_subgroup + abelian_group G +
assumes a_normal: "normal H \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
@@ -378,7 +378,7 @@
lemma (in abelian_subgroup) rcosets_add_eq:
"M \<in> a_rcosets H \<Longrightarrow> H <+> M = M"
- -- \<open>generalizes @{text subgroup_mult_id}\<close>
+ \<comment> \<open>generalizes \<open>subgroup_mult_id\<close>\<close>
by (rule normal.rcosets_mult_eq [OF a_normal,
folded set_add_def A_RCOSETS_def, simplified monoid_record_simps])
@@ -687,7 +687,7 @@
by (rule subgroup.rcos_module [OF a_subgroup a_group,
folded a_r_coset_def a_inv_def, simplified monoid_record_simps])
---"variant"
+\<comment>"variant"
lemma (in abelian_subgroup) a_rcos_module_minus:
assumes "ring G"
assumes carr: "x \<in> carrier G" "x' \<in> carrier G"
--- a/src/HOL/Algebra/Bij.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Algebra/Bij.thy Thu May 26 17:51:22 2016 +0200
@@ -10,7 +10,7 @@
definition
Bij :: "'a set \<Rightarrow> ('a \<Rightarrow> 'a) set"
- --\<open>Only extensional functions, since otherwise we get too many.\<close>
+ \<comment>\<open>Only extensional functions, since otherwise we get too many.\<close>
where "Bij S = extensional S \<inter> {f. bij_betw f S S}"
definition
--- a/src/HOL/Algebra/Congruence.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Algebra/Congruence.thy Thu May 26 17:51:22 2016 +0200
@@ -15,7 +15,7 @@
carrier :: "'a set"
-subsection \<open>Structure with Carrier and Equivalence Relation @{text eq}\<close>
+subsection \<open>Structure with Carrier and Equivalence Relation \<open>eq\<close>\<close>
record 'a eq_object = "'a partial_object" +
eq :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infixl ".=\<index>" 50)
--- a/src/HOL/Algebra/Coset.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Algebra/Coset.thy Thu May 26 17:51:22 2016 +0200
@@ -85,7 +85,7 @@
lemma (in group) coset_join2:
"\<lbrakk>x \<in> carrier G; subgroup H G; x\<in>H\<rbrakk> \<Longrightarrow> H #> x = H"
- --\<open>Alternative proof is to put @{term "x=\<one>"} in @{text repr_independence}.\<close>
+ \<comment>\<open>Alternative proof is to put @{term "x=\<one>"} in \<open>repr_independence\<close>.\<close>
by (force simp add: subgroup.m_closed r_coset_def solve_equation)
lemma (in monoid) r_coset_subset_G:
@@ -171,7 +171,7 @@
qed
qed
-text \<open>Step one for lemma @{text "rcos_module"}\<close>
+text \<open>Step one for lemma \<open>rcos_module\<close>\<close>
lemma (in subgroup) rcos_module_imp:
assumes "group G"
assumes xcarr: "x \<in> carrier G"
@@ -211,7 +211,7 @@
show "x' \<otimes> (inv x) \<in> H" by simp
qed
-text \<open>Step two for lemma @{text "rcos_module"}\<close>
+text \<open>Step two for lemma \<open>rcos_module\<close>\<close>
lemma (in subgroup) rcos_module_rev:
assumes "group G"
assumes carr: "x \<in> carrier G" "x' \<in> carrier G"
@@ -524,7 +524,7 @@
done
-subsubsection \<open>Set of Inverses of an @{text r_coset}.\<close>
+subsubsection \<open>Set of Inverses of an \<open>r_coset\<close>.\<close>
lemma (in normal) rcos_inv:
assumes x: "x \<in> carrier G"
@@ -549,7 +549,7 @@
qed
-subsubsection \<open>Theorems for @{text "<#>"} with @{text "#>"} or @{text "<#"}.\<close>
+subsubsection \<open>Theorems for \<open><#>\<close> with \<open>#>\<close> or \<open><#\<close>.\<close>
lemma (in group) setmult_rcos_assoc:
"\<lbrakk>H \<subseteq> carrier G; K \<subseteq> carrier G; x \<in> carrier G\<rbrakk>
@@ -584,7 +584,7 @@
by (simp add: rcos_mult_step1 rcos_mult_step2 rcos_mult_step3)
lemma (in normal) rcosets_mult_eq: "M \<in> rcosets H \<Longrightarrow> H <#> M = M"
- -- \<open>generalizes @{text subgroup_mult_id}\<close>
+ \<comment> \<open>generalizes \<open>subgroup_mult_id\<close>\<close>
by (auto simp add: RCOSETS_def subset
setmult_rcos_assoc subgroup_mult_id normal.axioms normal_axioms)
@@ -628,7 +628,7 @@
qed
qed
-text\<open>Equivalence classes of @{text rcong} correspond to left cosets.
+text\<open>Equivalence classes of \<open>rcong\<close> correspond to left cosets.
Was there a mistake in the definitions? I'd have expected them to
correspond to right cosets.\<close>
@@ -679,7 +679,7 @@
qed
-subsection \<open>Further lemmas for @{text "r_congruent"}\<close>
+subsection \<open>Further lemmas for \<open>r_congruent\<close>\<close>
text \<open>The relation is a congruence\<close>
@@ -780,7 +780,7 @@
apply (simp add: r_coset_subset_G [THEN finite_subset])
done
-text\<open>The next two lemmas support the proof of @{text card_cosets_equal}.\<close>
+text\<open>The next two lemmas support the proof of \<open>card_cosets_equal\<close>.\<close>
lemma (in group) inj_on_f:
"\<lbrakk>H \<subseteq> carrier G; a \<in> carrier G\<rbrakk> \<Longrightarrow> inj_on (\<lambda>y. y \<otimes> inv a) (H #> a)"
apply (rule inj_onI)
@@ -831,7 +831,7 @@
definition
FactGroup :: "[('a,'b) monoid_scheme, 'a set] \<Rightarrow> ('a set) monoid" (infixl "Mod" 65)
- --\<open>Actually defined for groups rather than monoids\<close>
+ \<comment>\<open>Actually defined for groups rather than monoids\<close>
where "FactGroup G H = \<lparr>carrier = rcosets\<^bsub>G\<^esub> H, mult = set_mult G, one = H\<rparr>"
lemma (in normal) setmult_closed:
@@ -897,7 +897,7 @@
definition
kernel :: "('a, 'm) monoid_scheme \<Rightarrow> ('b, 'n) monoid_scheme \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a set"
- --\<open>the kernel of a homomorphism\<close>
+ \<comment>\<open>the kernel of a homomorphism\<close>
where "kernel G H h = {x. x \<in> carrier G & h x = \<one>\<^bsub>H\<^esub>}"
lemma (in group_hom) subgroup_kernel: "subgroup (kernel G H h) G"
--- a/src/HOL/Algebra/Divisibility.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Algebra/Divisibility.thy Thu May 26 17:51:22 2016 +0200
@@ -2340,7 +2340,7 @@
qed
---"A version using @{const factors}, more complicated"
+\<comment>"A version using @{const factors}, more complicated"
lemma (in factorial_monoid) factors_irreducible_is_prime:
assumes pirr: "irreducible G p"
and pcarr: "p \<in> carrier G"
--- a/src/HOL/Algebra/FiniteProduct.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Algebra/FiniteProduct.thy Thu May 26 17:51:22 2016 +0200
@@ -12,10 +12,10 @@
subsubsection \<open>Inductive Definition of a Relation for Products over Sets\<close>
-text \<open>Instantiation of locale @{text LC} of theory @{text Finite_Set} is not
+text \<open>Instantiation of locale \<open>LC\<close> of theory \<open>Finite_Set\<close> is not
possible, because here we have explicit typing rules like
- @{text "x \<in> carrier G"}. We introduce an explicit argument for the domain
- @{text D}.\<close>
+ \<open>x \<in> carrier G\<close>. We introduce an explicit argument for the domain
+ \<open>D\<close>.\<close>
inductive_set
foldSetD :: "['a set, 'b => 'a => 'a, 'a] => ('b set * 'a) set"
@@ -111,7 +111,7 @@
apply (erule foldSetD.cases)
apply blast
apply clarify
- txt \<open>force simplification of @{text "card A < card (insert ...)"}.\<close>
+ txt \<open>force simplification of \<open>card A < card (insert ...)\<close>.\<close>
apply (erule rev_mp)
apply (simp add: less_Suc_eq_le)
apply (rule impI)
@@ -221,7 +221,7 @@
==> foldD D f e (A Un B) = foldD D f (foldD D f e B) A"
by (simp add: foldD_nest_Un_Int)
--- \<open>Delete rules to do with @{text foldSetD} relation.\<close>
+\<comment> \<open>Delete rules to do with \<open>foldSetD\<close> relation.\<close>
declare foldSetD_imp_finite [simp del]
empty_foldSetDE [rule del]
@@ -233,8 +233,8 @@
text \<open>Commutative Monoids\<close>
text \<open>
- We enter a more restrictive context, with @{text "f :: 'a => 'a => 'a"}
- instead of @{text "'b => 'a => 'a"}.
+ We enter a more restrictive context, with \<open>f :: 'a => 'a => 'a\<close>
+ instead of \<open>'b => 'a => 'a\<close>.
\<close>
locale ACeD =
@@ -299,7 +299,7 @@
("(3\<Otimes>__\<in>_. _)" [1000, 0, 51, 10] 10)
translations
"\<Otimes>\<^bsub>G\<^esub>i\<in>A. b" \<rightleftharpoons> "CONST finprod G (%i. b) A"
- -- \<open>Beware of argument permutation!\<close>
+ \<comment> \<open>Beware of argument permutation!\<close>
lemma (in comm_monoid) finprod_empty [simp]:
"finprod G f {} = \<one>"
@@ -367,7 +367,7 @@
"[| finite A; finite B; g \<in> A \<rightarrow> carrier G; g \<in> B \<rightarrow> carrier G |] ==>
finprod G g (A Un B) \<otimes> finprod G g (A Int B) =
finprod G g A \<otimes> finprod G g B"
--- \<open>The reversed orientation looks more natural, but LOOPS as a simprule!\<close>
+\<comment> \<open>The reversed orientation looks more natural, but LOOPS as a simprule!\<close>
proof (induct set: finite)
case empty then show ?case by simp
next
@@ -451,7 +451,7 @@
by (rule finprod_cong') (auto simp add: simp_implies_def)
text \<open>Usually, if this rule causes a failed congruence proof error,
- the reason is that the premise @{text "g \<in> B \<rightarrow> carrier G"} cannot be shown.
+ the reason is that the premise \<open>g \<in> B \<rightarrow> carrier G\<close> cannot be shown.
Adding @{thm [source] Pi_def} to the simpset is often useful.
For this reason, @{thm [source] finprod_cong}
is not added to the simpset by default.
--- a/src/HOL/Algebra/Group.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Algebra/Group.thy Thu May 26 17:51:22 2016 +0200
@@ -26,7 +26,7 @@
definition
Units :: "_ => 'a set"
- --\<open>The set of invertible elements\<close>
+ \<comment>\<open>The set of invertible elements\<close>
where "Units G = {y. y \<in> carrier G & (\<exists>x \<in> carrier G. x \<otimes>\<^bsub>G\<^esub> y = \<one>\<^bsub>G\<^esub> & y \<otimes>\<^bsub>G\<^esub> x = \<one>\<^bsub>G\<^esub>)}"
consts
@@ -98,7 +98,7 @@
moreover from x y xinv yinv have "x \<otimes> (y \<otimes> y') \<otimes> x' = \<one>" by simp
moreover note x y
ultimately show ?thesis unfolding Units_def
- -- "Must avoid premature use of @{text hyp_subst_tac}."
+ \<comment> "Must avoid premature use of \<open>hyp_subst_tac\<close>."
apply (rule_tac CollectI)
apply (rule)
apply (fast)
@@ -486,8 +486,8 @@
text \<open>
Since @{term H} is nonempty, it contains some element @{term x}. Since
- it is closed under inverse, it contains @{text "inv x"}. Since
- it is closed under product, it contains @{text "x \<otimes> inv x = \<one>"}.
+ it is closed under inverse, it contains \<open>inv x\<close>. Since
+ it is closed under product, it contains \<open>x \<otimes> inv x = \<one>\<close>.
\<close>
lemma (in group) one_in_subset:
--- a/src/HOL/Algebra/Ideal.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Algebra/Ideal.thy Thu May 26 17:51:22 2016 +0200
@@ -592,14 +592,14 @@
subsection \<open>Properties of Principal Ideals\<close>
-text \<open>@{text "\<zero>"} generates the zero ideal\<close>
+text \<open>\<open>\<zero>\<close> generates the zero ideal\<close>
lemma (in ring) zero_genideal: "Idl {\<zero>} = {\<zero>}"
apply rule
apply (simp add: genideal_minimal zeroideal)
apply (fast intro!: genideal_self)
done
-text \<open>@{text "\<one>"} generates the unit ideal\<close>
+text \<open>\<open>\<one>\<close> generates the unit ideal\<close>
lemma (in ring) one_genideal: "Idl {\<one>} = carrier R"
proof -
have "\<one> \<in> Idl {\<one>}"
@@ -684,7 +684,7 @@
then show thesis using primeidealCD [OF R.is_cring notprime] by blast
qed
-text \<open>If @{text "{\<zero>}"} is a prime ideal of a commutative ring, the ring is a domain\<close>
+text \<open>If \<open>{\<zero>}\<close> is a prime ideal of a commutative ring, the ring is a domain\<close>
lemma (in cring) zeroprimeideal_domainI:
assumes pi: "primeideal {\<zero>} R"
shows "domain R"
@@ -828,7 +828,7 @@
subsection \<open>Derived Theorems\<close>
---"A non-zero cring that has only the two trivial ideals is a field"
+\<comment>"A non-zero cring that has only the two trivial ideals is a field"
lemma (in cring) trivialideals_fieldI:
assumes carrnzero: "carrier R \<noteq> {\<zero>}"
and haveideals: "{I. ideal I R} = {{\<zero>}, carrier R}"
@@ -921,7 +921,7 @@
qed
qed (simp add: zeroideal oneideal)
---"Jacobson Theorem 2.2"
+\<comment>"Jacobson Theorem 2.2"
lemma (in cring) trivialideals_eq_field:
assumes carrnzero: "carrier R \<noteq> {\<zero>}"
shows "({I. ideal I R} = {{\<zero>}, carrier R}) = field R"
--- a/src/HOL/Algebra/IntRing.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Algebra/IntRing.thy Thu May 26 17:51:22 2016 +0200
@@ -20,7 +20,7 @@
done
-subsection \<open>@{text "\<Z>"}: The Set of Integers as Algebraic Structure\<close>
+subsection \<open>\<open>\<Z>\<close>: The Set of Integers as Algebraic Structure\<close>
abbreviation int_ring :: "int ring" ("\<Z>")
where "int_ring \<equiv> \<lparr>carrier = UNIV, mult = op *, one = 1, zero = 0, add = op +\<rparr>"
@@ -59,14 +59,14 @@
and "one \<Z> = 1"
and "pow \<Z> x n = x^n"
proof -
- -- "Specification"
+ \<comment> "Specification"
show "monoid \<Z>" by standard auto
then interpret int: monoid \<Z> .
- -- "Carrier"
+ \<comment> "Carrier"
show "carrier \<Z> = UNIV" by simp
- -- "Operations"
+ \<comment> "Operations"
{ fix x y show "mult \<Z> x y = x * y" by simp }
show "one \<Z> = 1" by simp
show "pow \<Z> x n = x^n" by (induct n) simp_all
@@ -75,11 +75,11 @@
interpretation int: comm_monoid \<Z>
rewrites "finprod \<Z> f A = setprod f A"
proof -
- -- "Specification"
+ \<comment> "Specification"
show "comm_monoid \<Z>" by standard auto
then interpret int: comm_monoid \<Z> .
- -- "Operations"
+ \<comment> "Operations"
{ fix x y have "mult \<Z> x y = x * y" by simp }
note mult = this
have one: "one \<Z> = 1" by simp
@@ -93,14 +93,14 @@
and int_add_eq: "add \<Z> x y = x + y"
and int_finsum_eq: "finsum \<Z> f A = setsum f A"
proof -
- -- "Specification"
+ \<comment> "Specification"
show "abelian_monoid \<Z>" by standard auto
then interpret int: abelian_monoid \<Z> .
- -- "Carrier"
+ \<comment> "Carrier"
show "carrier \<Z> = UNIV" by simp
- -- "Operations"
+ \<comment> "Operations"
{ fix x y show "add \<Z> x y = x + y" by simp }
note add = this
show zero: "zero \<Z> = 0"
@@ -121,7 +121,7 @@
and int_a_inv_eq: "a_inv \<Z> x = - x"
and int_a_minus_eq: "a_minus \<Z> x y = x - y"
proof -
- -- "Specification"
+ \<comment> "Specification"
show "abelian_group \<Z>"
proof (rule abelian_groupI)
fix x
@@ -130,7 +130,7 @@
by simp arith
qed auto
then interpret int: abelian_group \<Z> .
- -- "Operations"
+ \<comment> "Operations"
{ fix x y have "add \<Z> x y = x + y" by simp }
note add = this
have zero: "zero \<Z> = 0" by simp
@@ -218,7 +218,7 @@
by standard clarsimp
-subsection \<open>Generated Ideals of @{text "\<Z>"}\<close>
+subsection \<open>Generated Ideals of \<open>\<Z>\<close>\<close>
lemma int_Idl: "Idl\<^bsub>\<Z>\<^esub> {a} = {x * a | x. True}"
apply (subst int.cgenideal_eq_genideal[symmetric]) apply simp
--- a/src/HOL/Algebra/Lattice.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Algebra/Lattice.thy Thu May 26 17:51:22 2016 +0200
@@ -507,7 +507,7 @@
unfolding sup_def
by (rule someI2) (auto intro: sup_of_singletonI)
-text \<open>Condition on @{text A}: supremum exists.\<close>
+text \<open>Condition on \<open>A\<close>: supremum exists.\<close>
lemma (in weak_upper_semilattice) sup_insertI:
"[| !!s. least L s (Upper L (insert x A)) ==> P s;
@@ -638,7 +638,7 @@
assumes L: "x \<in> carrier L" "y \<in> carrier L" "z \<in> carrier L"
shows "x \<squnion> (y \<squnion> z) .= \<Squnion>{x, y, z}"
proof (rule finite_sup_insertI)
- -- \<open>The textbook argument in Jacobson I, p 457\<close>
+ \<comment> \<open>The textbook argument in Jacobson I, p 457\<close>
fix s
assume sup: "least L s (Upper L {x, y, z})"
show "x \<squnion> (y \<squnion> z) .= s"
@@ -652,7 +652,7 @@
qed (simp_all add: L least_closed [OF sup])
qed (simp_all add: L)
-text \<open>Commutativity holds for @{text "="}.\<close>
+text \<open>Commutativity holds for \<open>=\<close>.\<close>
lemma join_comm:
fixes L (structure)
@@ -747,7 +747,7 @@
unfolding inf_def
by (rule someI2) (auto intro: inf_of_singletonI)
-text \<open>Condition on @{text A}: infimum exists.\<close>
+text \<open>Condition on \<open>A\<close>: infimum exists.\<close>
lemma (in weak_lower_semilattice) inf_insertI:
"[| !!i. greatest L i (Lower L (insert x A)) ==> P i;
@@ -1090,7 +1090,7 @@
(* TODO: prove dual version *)
-subsection \<open>Orders and Lattices where @{text eq} is the Equality\<close>
+subsection \<open>Orders and Lattices where \<open>eq\<close> is the Equality\<close>
locale partial_order = weak_partial_order +
assumes eq_is_equal: "op .= = op ="
@@ -1172,7 +1172,7 @@
"x \<in> carrier L ==> \<Sqinter>{x} = x"
using weak_inf_of_singleton unfolding eq_is_equal .
-text \<open>Condition on @{text A}: infimum exists.\<close>
+text \<open>Condition on \<open>A\<close>: infimum exists.\<close>
lemma (in lower_semilattice) meet_assoc_lemma:
assumes L: "x \<in> carrier L" "y \<in> carrier L" "z \<in> carrier L"
--- a/src/HOL/Algebra/QuotRing.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Algebra/QuotRing.thy Thu May 26 17:51:22 2016 +0200
@@ -84,28 +84,28 @@
text \<open>The quotient is a ring\<close>
lemma (in ideal) quotient_is_ring: "ring (R Quot I)"
apply (rule ringI)
- --\<open>abelian group\<close>
+ \<comment>\<open>abelian group\<close>
apply (rule comm_group_abelian_groupI)
apply (simp add: FactRing_def)
apply (rule a_factorgroup_is_comm_group[unfolded A_FactGroup_def'])
- --\<open>mult monoid\<close>
+ \<comment>\<open>mult monoid\<close>
apply (rule monoidI)
apply (simp_all add: FactRing_def A_RCOSETS_def RCOSETS_def
a_r_coset_def[symmetric])
- --\<open>mult closed\<close>
+ \<comment>\<open>mult closed\<close>
apply (clarify)
apply (simp add: rcoset_mult_add, fast)
- --\<open>mult @{text one_closed}\<close>
+ \<comment>\<open>mult \<open>one_closed\<close>\<close>
apply force
- --\<open>mult assoc\<close>
+ \<comment>\<open>mult assoc\<close>
apply clarify
apply (simp add: rcoset_mult_add m_assoc)
- --\<open>mult one\<close>
+ \<comment>\<open>mult one\<close>
apply clarify
apply (simp add: rcoset_mult_add)
apply clarify
apply (simp add: rcoset_mult_add)
- --\<open>distr\<close>
+ \<comment>\<open>distr\<close>
apply clarify
apply (simp add: rcoset_mult_add a_rcos_sum l_distr)
apply clarify
@@ -225,7 +225,7 @@
apply (simp add: FactRing_def A_RCOSETS_defs a_r_coset_def[symmetric], clarsimp)
apply (simp add: rcoset_mult_add) defer 1
proof (rule ccontr, simp)
- --\<open>Quotient is not empty\<close>
+ \<comment>\<open>Quotient is not empty\<close>
assume "\<zero>\<^bsub>R Quot I\<^esub> = \<one>\<^bsub>R Quot I\<^esub>"
then have II1: "I = I +> \<one>" by (simp add: FactRing_def)
from a_rcos_self[OF one_closed] have "\<one> \<in> I"
@@ -233,11 +233,11 @@
then have "I = carrier R" by (rule one_imp_carrier)
with I_notcarr show False by simp
next
- --\<open>Existence of Inverse\<close>
+ \<comment>\<open>Existence of Inverse\<close>
fix a
assume IanI: "I +> a \<noteq> I" and acarr: "a \<in> carrier R"
- --\<open>Helper ideal @{text "J"}\<close>
+ \<comment>\<open>Helper ideal \<open>J\<close>\<close>
define J :: "'a set" where "J = (carrier R #> a) <+> I"
have idealJ: "ideal J R"
apply (unfold J_def, rule add_ideals)
@@ -245,7 +245,7 @@
apply (rule is_ideal)
done
- --\<open>Showing @{term "J"} not smaller than @{term "I"}\<close>
+ \<comment>\<open>Showing @{term "J"} not smaller than @{term "I"}\<close>
have IinJ: "I \<subseteq> J"
proof (rule, simp add: J_def r_coset_def set_add_defs)
fix x
@@ -256,7 +256,7 @@
with Zcarr and xI show "\<exists>xa\<in>carrier R. \<exists>k\<in>I. x = xa \<otimes> a \<oplus> k" by fast
qed
- --\<open>Showing @{term "J \<noteq> I"}\<close>
+ \<comment>\<open>Showing @{term "J \<noteq> I"}\<close>
have anI: "a \<notin> I"
proof (rule ccontr, simp)
assume "a \<in> I"
@@ -274,7 +274,7 @@
from aJ and anI have JnI: "J \<noteq> I" by fast
- --\<open>Deducing @{term "J = carrier R"} because @{term "I"} is maximal\<close>
+ \<comment>\<open>Deducing @{term "J = carrier R"} because @{term "I"} is maximal\<close>
from idealJ and IinJ have "J = I \<or> J = carrier R"
proof (rule I_maximal, unfold J_def)
have "carrier R #> a \<subseteq> carrier R"
@@ -285,7 +285,7 @@
with JnI have Jcarr: "J = carrier R" by simp
- --\<open>Calculating an inverse for @{term "a"}\<close>
+ \<comment>\<open>Calculating an inverse for @{term "a"}\<close>
from one_closed[folded Jcarr]
have "\<exists>r\<in>carrier R. \<exists>i\<in>I. \<one> = r \<otimes> a \<oplus> i"
by (simp add: J_def r_coset_def set_add_defs)
@@ -294,7 +294,7 @@
from one and rcarr and acarr and iI[THEN a_Hcarr]
have rai1: "a \<otimes> r = \<ominus>i \<oplus> \<one>" by algebra
- --\<open>Lifting to cosets\<close>
+ \<comment>\<open>Lifting to cosets\<close>
from iI have "\<ominus>i \<oplus> \<one> \<in> I +> \<one>"
by (intro a_rcosI, simp, intro a_subset, simp)
with rai1 have "a \<otimes> r \<in> I +> \<one>" by simp
--- a/src/HOL/Algebra/Ring.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Algebra/Ring.thy Thu May 26 17:51:22 2016 +0200
@@ -39,7 +39,7 @@
("(3\<Oplus>__\<in>_. _)" [1000, 0, 51, 10] 10)
translations
"\<Oplus>\<^bsub>G\<^esub>i\<in>A. b" \<rightleftharpoons> "CONST finsum G (%i. b) A"
- -- \<open>Beware of argument permutation!\<close>
+ \<comment> \<open>Beware of argument permutation!\<close>
locale abelian_group = abelian_monoid +
@@ -141,7 +141,7 @@
lemmas finsum_cong = add.finprod_cong
text \<open>Usually, if this rule causes a failed congruence proof error,
- the reason is that the premise @{text "g \<in> B \<rightarrow> carrier G"} cannot be shown.
+ the reason is that the premise \<open>g \<in> B \<rightarrow> carrier G\<close> cannot be shown.
Adding @{thm [source] Pi_def} to the simpset is often useful.\<close>
lemmas finsum_reindex = add.finprod_reindex
@@ -206,7 +206,7 @@
lemmas (in abelian_group) minus_add = add.inv_mult
-text \<open>Derive an @{text "abelian_group"} from a @{text "comm_group"}\<close>
+text \<open>Derive an \<open>abelian_group\<close> from a \<open>comm_group\<close>\<close>
lemma comm_group_abelian_groupI:
fixes G (structure)
@@ -283,7 +283,7 @@
shows "cring R"
proof (intro cring.intro ring.intro)
show "ring_axioms R"
- -- \<open>Right-distributivity follows from left-distributivity and
+ \<comment> \<open>Right-distributivity follows from left-distributivity and
commutativity.\<close>
proof (rule ring_axioms.intro)
fix x y z
--- a/src/HOL/Algebra/RingHom.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Algebra/RingHom.thy Thu May 26 17:51:22 2016 +0200
@@ -8,7 +8,7 @@
section \<open>Homomorphisms of Non-Commutative Rings\<close>
-text \<open>Lifting existing lemmas in a @{text ring_hom_ring} locale\<close>
+text \<open>Lifting existing lemmas in a \<open>ring_hom_ring\<close> locale\<close>
locale ring_hom_ring = R?: ring R + S?: ring S
for R (structure) and S (structure) +
fixes h
@@ -102,7 +102,7 @@
subsection \<open>The Kernel of a Ring Homomorphism\<close>
---"the kernel of a ring homomorphism is an ideal"
+\<comment>"the kernel of a ring homomorphism is an ideal"
lemma (in ring_hom_ring) kernel_is_ideal:
shows "ideal (a_kernel R S h) R"
apply (rule idealI)
--- a/src/HOL/Algebra/Sylow.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Algebra/Sylow.thy Thu May 26 17:51:22 2016 +0200
@@ -104,7 +104,7 @@
end
-subsection\<open>Discharging the Assumptions of @{text sylow_central}\<close>
+subsection\<open>Discharging the Assumptions of \<open>sylow_central\<close>\<close>
context sylow
begin
@@ -265,7 +265,7 @@
intro!: M1_in_M in_quotient_imp_closed [OF RelM_equiv M_in_quot _ M1_RelM_rcosetGM1g])
done
-text\<open>close to a duplicate of @{text inj_M_GmodH}\<close>
+text\<open>close to a duplicate of \<open>inj_M_GmodH\<close>\<close>
lemma (in sylow_central) inj_GmodH_M:
"\<exists>g \<in> rcosets H\<rightarrow>M. inj_on g (rcosets H)"
apply (rule bexI)
--- a/src/HOL/Algebra/UnivPoly.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Algebra/UnivPoly.thy Thu May 26 17:51:22 2016 +0200
@@ -14,7 +14,7 @@
text \<open>
Polynomials are formalised as modules with additional operations for
extracting coefficients from polynomials and for obtaining monomials
- from coefficients and exponents (record @{text "up_ring"}). The
+ from coefficients and exponents (record \<open>up_ring\<close>). The
carrier set is a set of bounded functions from Nat to the
coefficient domain. Bounded means that these functions return zero
above a certain bound (the degree). There is a chapter on the
@@ -754,7 +754,7 @@
assumes non_zero: "n ~= 0 ==> coeff P p n ~= \<zero>"
and R: "p \<in> carrier P"
shows "n <= deg R p"
--- \<open>Logically, this is a slightly stronger version of
+\<comment> \<open>Logically, this is a slightly stronger version of
@{thm [source] deg_aboveD}\<close>
proof (cases "n=0")
case True then show ?thesis by simp
@@ -864,7 +864,7 @@
qed
text\<open>The following lemma is later \emph{overwritten} by the most
- specific one for domains, @{text deg_smult}.\<close>
+ specific one for domains, \<open>deg_smult\<close>.\<close>
lemma deg_smult_ring [simp]:
"[| a \<in> carrier R; p \<in> carrier P |] ==>
@@ -1202,7 +1202,7 @@
defines Eval_def: "Eval == eval R S h s"
text\<open>JE: I have moved the following lemma from Ring.thy and lifted then to the locale @{term ring_hom_ring} from @{term ring_hom_cring}.\<close>
-text\<open>JE: I was considering using it in @{text eval_ring_hom}, but that property does not hold for non commutative rings, so
+text\<open>JE: I was considering using it in \<open>eval_ring_hom\<close>, but that property does not hold for non commutative rings, so
maybe it is not that necessary.\<close>
lemma (in ring_hom_ring) hom_finsum [simp]:
@@ -1295,8 +1295,8 @@
qed
text \<open>
- The following lemma could be proved in @{text UP_cring} with the additional
- assumption that @{text h} is closed.\<close>
+ The following lemma could be proved in \<open>UP_cring\<close> with the additional
+ assumption that \<open>h\<close> is closed.\<close>
lemma (in UP_pre_univ_prop) eval_const:
"[| s \<in> carrier S; r \<in> carrier R |] ==> eval R S h s (monom P r 0) = h r"
--- a/src/HOL/Cardinals/Bounded_Set.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Cardinals/Bounded_Set.thy Thu May 26 17:51:22 2016 +0200
@@ -61,9 +61,9 @@
hence *: "set_bset R' = ?L'" unfolding R'_def by (intro set_bset_to_set_bset)
show ?R unfolding Grp_def relcompp.simps conversep.simps
proof (intro CollectI case_prodI exI[of _ a] exI[of _ b] exI[of _ R'] conjI refl)
- from * show "a = map_bset fst R'" using conjunct1[OF `?L`]
+ from * show "a = map_bset fst R'" using conjunct1[OF \<open>?L\<close>]
by (transfer, auto simp add: image_def Int_def split: prod.splits)
- from * show "b = map_bset snd R'" using conjunct2[OF `?L`]
+ from * show "b = map_bset snd R'" using conjunct2[OF \<open>?L\<close>]
by (transfer, auto simp add: image_def Int_def split: prod.splits)
qed (auto simp add: *)
next
--- a/src/HOL/Cardinals/Cardinal_Arithmetic.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Cardinals/Cardinal_Arithmetic.thy Thu May 26 17:51:22 2016 +0200
@@ -5,13 +5,13 @@
Cardinal arithmetic.
*)
-section {* Cardinal Arithmetic *}
+section \<open>Cardinal Arithmetic\<close>
theory Cardinal_Arithmetic
imports BNF_Cardinal_Arithmetic Cardinal_Order_Relation
begin
-subsection {* Binary sum *}
+subsection \<open>Binary sum\<close>
lemma csum_Cnotzero2:
"Cnotzero r2 \<Longrightarrow> Cnotzero (r1 +c r2)"
@@ -41,7 +41,7 @@
by (rule ordIso_transitive[OF ordIso_symmetric[OF card_of_Plus_empty2] card_of_Field_ordIso])
-subsection {* Product *}
+subsection \<open>Product\<close>
lemma Times_cprod: "|A \<times> B| =o |A| *c |B|"
by (simp only: cprod_def Field_card_of card_of_refl)
@@ -71,7 +71,7 @@
unfolding cprod_def by (metis Card_order_Times1 czeroI)
-subsection {* Exponentiation *}
+subsection \<open>Exponentiation\<close>
lemma cexp_czero: "r ^c czero =o cone"
unfolding cexp_def czero_def Field_card_of Func_empty by (rule single_cone)
@@ -269,7 +269,7 @@
unfolding cexp_def cprod_def csum_def Field_card_of by (rule card_of_Func_Plus)
-subsection {* Powerset *}
+subsection \<open>Powerset\<close>
definition cpow where "cpow r = |Pow (Field r)|"
@@ -291,7 +291,7 @@
lemma cpow_cexp_ctwo: "cpow r =o ctwo ^c r"
unfolding cpow_def ctwo_def cexp_def Field_card_of by (rule card_of_Pow_Func)
-subsection {* Inverse image *}
+subsection \<open>Inverse image\<close>
lemma vimage_ordLeq:
assumes "|A| \<le>o k" and "\<forall> a \<in> A. |vimage f {a}| \<le>o k" and "Cinfinite k"
@@ -303,7 +303,7 @@
finally show ?thesis .
qed
-subsection {* Maximum *}
+subsection \<open>Maximum\<close>
definition cmax where
"cmax r s =
@@ -372,7 +372,7 @@
case False
hence "s \<le>o r" by (metis ordLeq_total r s card_order_on_def)
hence "cmax r s =o r" by (metis cmax1 r s)
- also have "r =o r +c s" by (metis `s \<le>o r` csum_absorb1 ordIso_symmetric r)
+ also have "r =o r +c s" by (metis \<open>s \<le>o r\<close> csum_absorb1 ordIso_symmetric r)
finally show ?thesis .
qed
@@ -386,7 +386,7 @@
case False
hence "s \<le>o r" by (metis ordLeq_total r s card_order_on_def)
hence "cmax r s =o r" by (metis cmax1 r s)
- also have "r =o r *c s" by (metis Cinfinite_Cnotzero `s \<le>o r` cprod_infinite1' ordIso_symmetric r s)
+ also have "r =o r *c s" by (metis Cinfinite_Cnotzero \<open>s \<le>o r\<close> cprod_infinite1' ordIso_symmetric r s)
finally show ?thesis .
qed
--- a/src/HOL/Cardinals/Cardinal_Order_Relation.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Cardinals/Cardinal_Order_Relation.thy Thu May 26 17:51:22 2016 +0200
@@ -5,7 +5,7 @@
Cardinal-order relations.
*)
-section {* Cardinal-Order Relations *}
+section \<open>Cardinal-Order Relations\<close>
theory Cardinal_Order_Relation
imports BNF_Cardinal_Order_Relation Wellorder_Constructions
@@ -69,7 +69,7 @@
curr_in[intro, simp]
-subsection {* Cardinal of a set *}
+subsection \<open>Cardinal of a set\<close>
lemma card_of_inj_rel: assumes INJ: "!! x y y'. \<lbrakk>(x,y) : R; (x,y') : R\<rbrakk> \<Longrightarrow> y = y'"
shows "|{y. EX x. (x,y) : R}| <=o |{x. EX y. (x,y) : R}|"
@@ -132,7 +132,7 @@
qed
-subsection {* Cardinals versus set operations on arbitrary sets *}
+subsection \<open>Cardinals versus set operations on arbitrary sets\<close>
lemma card_of_set_type[simp]: "|UNIV::'a set| <o |UNIV::'a set set|"
using card_of_Pow[of "UNIV::'a set"] by simp
@@ -659,7 +659,7 @@
by (metis assms ordIso_iff_ordLeq ordLeq_finite_Field)
-subsection {* Cardinals versus set operations involving infinite sets *}
+subsection \<open>Cardinals versus set operations involving infinite sets\<close>
lemma finite_iff_cardOf_nat:
"finite A = ( |A| <o |UNIV :: nat set| )"
@@ -704,10 +704,10 @@
qed
-subsection {* Cardinals versus lists *}
+subsection \<open>Cardinals versus lists\<close>
-text{* The next is an auxiliary operator, which shall be used for inductive
-proofs of facts concerning the cardinality of @{text "List"} : *}
+text\<open>The next is an auxiliary operator, which shall be used for inductive
+proofs of facts concerning the cardinality of \<open>List\<close> :\<close>
definition nlists :: "'a set \<Rightarrow> nat \<Rightarrow> 'a list set"
where "nlists A n \<equiv> {l. set l \<le> A \<and> length l = n}"
@@ -886,7 +886,7 @@
using lists_UNIV by auto
-subsection {* Cardinals versus the set-of-finite-sets operator *}
+subsection \<open>Cardinals versus the set-of-finite-sets operator\<close>
definition Fpow :: "'a set \<Rightarrow> 'a set set"
where "Fpow A \<equiv> {X. X \<le> A \<and> finite X}"
@@ -999,9 +999,9 @@
using assms card_of_Fpow_infinite card_of_ordIso by blast
-subsection {* The cardinal $\omega$ and the finite cardinals *}
+subsection \<open>The cardinal $\omega$ and the finite cardinals\<close>
-subsubsection {* First as well-orders *}
+subsubsection \<open>First as well-orders\<close>
lemma Field_natLess: "Field natLess = (UNIV::nat set)"
by(unfold Field_def natLess_def, auto)
@@ -1111,7 +1111,7 @@
qed
-subsubsection {* Then as cardinals *}
+subsubsection \<open>Then as cardinals\<close>
lemma ordIso_natLeq_infinite1:
"|A| =o natLeq \<Longrightarrow> \<not>finite A"
@@ -1179,7 +1179,7 @@
qed
-subsubsection {* "Backward compatibility" with the numeric cardinal operator for finite sets *}
+subsubsection \<open>"Backward compatibility" with the numeric cardinal operator for finite sets\<close>
lemma finite_card_of_iff_card2:
assumes FIN: "finite A" and FIN': "finite B"
@@ -1221,7 +1221,7 @@
using Field_natLeq_on card_atLeastLessThan by auto
-subsection {* The successor of a cardinal *}
+subsection \<open>The successor of a cardinal\<close>
lemma embed_implies_ordIso_Restr:
assumes WELL: "Well_order r" and WELL': "Well_order r'" and EMB: "embed r' r f"
@@ -1321,7 +1321,7 @@
ordLeq_transitive by metis
-subsection {* Others *}
+subsection \<open>Others\<close>
lemma under_mono[simp]:
assumes "Well_order r" and "(i,j) \<in> r"
@@ -1575,7 +1575,7 @@
using ordLeq_Func[OF B] by (metis A card_of_ordLeq_finite)
-subsection {* Infinite cardinals are limit ordinals *}
+subsection \<open>Infinite cardinals are limit ordinals\<close>
lemma card_order_infinite_isLimOrd:
assumes c: "Card_order r" and i: "\<not>finite (Field r)"
@@ -1683,7 +1683,7 @@
qed
-subsection {* Regular vs. stable cardinals *}
+subsection \<open>Regular vs. stable cardinals\<close>
definition stable :: "'a rel \<Rightarrow> bool"
where
@@ -1851,9 +1851,9 @@
lemma stable_nat: "stable |UNIV::nat set|"
using stable_natLeq card_of_nat stable_ordIso by auto
-text{* Below, the type of "A" is not important -- we just had to choose an appropriate
+text\<open>Below, the type of "A" is not important -- we just had to choose an appropriate
type to make "A" possible. What is important is that arbitrarily large
- infinite sets of stable cardinality exist. *}
+ infinite sets of stable cardinality exist.\<close>
lemma infinite_stable_exists:
assumes CARD: "\<forall>r \<in> R. Card_order (r::'a rel)"
--- a/src/HOL/Cardinals/Cardinals.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Cardinals/Cardinals.thy Thu May 26 17:51:22 2016 +0200
@@ -6,7 +6,7 @@
Theory of ordinals and cardinals.
*)
-section {* Theory of Ordinals and Cardinals *}
+section \<open>Theory of Ordinals and Cardinals\<close>
theory Cardinals
imports Ordinal_Arithmetic Cardinal_Arithmetic Wellorder_Extension
--- a/src/HOL/Cardinals/Fun_More.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Cardinals/Fun_More.thy Thu May 26 17:51:22 2016 +0200
@@ -5,13 +5,13 @@
More on injections, bijections and inverses.
*)
-section {* More on Injections, Bijections and Inverses *}
+section \<open>More on Injections, Bijections and Inverses\<close>
theory Fun_More
imports Main
begin
-subsection {* Purely functional properties *}
+subsection \<open>Purely functional properties\<close>
(* unused *)
(*1*)lemma notIn_Un_bij_betw2:
@@ -56,7 +56,7 @@
qed
-subsection {* Properties involving finite and infinite sets *}
+subsection \<open>Properties involving finite and infinite sets\<close>
(*3*)lemma inj_on_image_Pow:
assumes "inj_on f A"
@@ -128,7 +128,7 @@
qed
-subsection {* Properties involving Hilbert choice *}
+subsection \<open>Properties involving Hilbert choice\<close>
(*1*)lemma bij_betw_inv_into_LEFT:
assumes BIJ: "bij_betw f A A'" and SUB: "B \<le> A"
@@ -142,7 +142,7 @@
using assms bij_betw_inv_into_LEFT[of f A A' B] by fast
-subsection {* Other facts *}
+subsection \<open>Other facts\<close>
(*3*)lemma atLeastLessThan_injective:
assumes "{0 ..< m::nat} = {0 ..< n}"
--- a/src/HOL/Cardinals/Order_Relation_More.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Cardinals/Order_Relation_More.thy Thu May 26 17:51:22 2016 +0200
@@ -5,13 +5,13 @@
Basics on order-like relations.
*)
-section {* Basics on Order-Like Relations *}
+section \<open>Basics on Order-Like Relations\<close>
theory Order_Relation_More
imports Main
begin
-subsection {* The upper and lower bounds operators *}
+subsection \<open>The upper and lower bounds operators\<close>
lemma aboveS_subset_above: "aboveS r a \<le> above r a"
by(auto simp add: aboveS_def above_def)
@@ -572,7 +572,7 @@
qed
-subsection {* Properties depending on more than one relation *}
+subsection \<open>Properties depending on more than one relation\<close>
lemma under_incr2:
"r \<le> r' \<Longrightarrow> under r a \<le> under r' a"
--- a/src/HOL/Cardinals/Order_Union.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Cardinals/Order_Union.thy Thu May 26 17:51:22 2016 +0200
@@ -4,7 +4,7 @@
The ordinal-like sum of two orders with disjoint fields
*)
-section {* Order Union *}
+section \<open>Order Union\<close>
theory Order_Union
imports Order_Relation
--- a/src/HOL/Cardinals/Ordinal_Arithmetic.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Cardinals/Ordinal_Arithmetic.thy Thu May 26 17:51:22 2016 +0200
@@ -5,7 +5,7 @@
Ordinal arithmetic.
*)
-section {* Ordinal Arithmetic *}
+section \<open>Ordinal Arithmetic\<close>
theory Ordinal_Arithmetic
imports Wellorder_Constructions
@@ -252,7 +252,7 @@
fix B
assume *: "B \<subseteq> Field {((x,y1), (x,y2)) . x \<in> A \<and> (y1, y2) \<in> r}" and "B \<noteq> {}"
from image_mono[OF *, of snd] have "snd ` B \<subseteq> Field r" unfolding Field_def by force
- with `B \<noteq> {}` obtain x where x: "x \<in> snd ` B" "\<forall>x'\<in>snd ` B. (x', x) \<notin> r"
+ with \<open>B \<noteq> {}\<close> obtain x where x: "x \<in> snd ` B" "\<forall>x'\<in>snd ` B. (x', x) \<notin> r"
using spec[OF assms[unfolded wf_eq_minimal2], of "snd ` B"] by auto
then obtain a where "(a, x) \<in> B" by auto
moreover
@@ -271,7 +271,7 @@
fix B
assume *: "B \<subseteq> Field {((x1, y), (x2, y)). (x1, x2) \<in> r \<and> y \<in> A}" and "B \<noteq> {}"
from image_mono[OF *, of fst] have "fst ` B \<subseteq> Field r" unfolding Field_def by force
- with `B \<noteq> {}` obtain x where x: "x \<in> fst ` B" "\<forall>x'\<in>fst ` B. (x', x) \<notin> r"
+ with \<open>B \<noteq> {}\<close> obtain x where x: "x \<in> fst ` B" "\<forall>x'\<in>fst ` B. (x', x) \<notin> r"
using spec[OF assms[unfolded wf_eq_minimal2], of "fst ` B"] by auto
then obtain a where "(x, a) \<in> B" by auto
moreover
@@ -364,9 +364,9 @@
(f o Inl) \<in> fin_support z A \<and> (f o Inr) \<in> fin_support z B" (is "?L \<longleftrightarrow> ?R1 \<and> ?R2")
proof safe
assume ?L
- from `?L` show ?R1 unfolding fin_support_def support_def
+ from \<open>?L\<close> show ?R1 unfolding fin_support_def support_def
by (fastforce simp: image_iff elim: finite_surj[of _ _ "case_sum id undefined"])
- from `?L` show ?R2 unfolding fin_support_def support_def
+ from \<open>?L\<close> show ?R2 unfolding fin_support_def support_def
by (fastforce simp: image_iff elim: finite_surj[of _ _ "case_sum undefined id"])
next
assume ?R1 ?R2
@@ -618,7 +618,7 @@
show ?thesis
proof (cases "?fg = ?gh \<longrightarrow> f ?fg \<noteq> h ?gh")
case True
- show ?thesis using max_fun_diff_max2[of f g h, OF True] * `f \<noteq> h` max_fun_diff_in
+ show ?thesis using max_fun_diff_max2[of f g h, OF True] * \<open>f \<noteq> h\<close> max_fun_diff_in
r.max2_iff[OF FINFUNCD FINFUNCD] r.max2_equals1[OF FINFUNCD FINFUNCD] max_fun_diff_le_eq
s.in_notinI[OF disjI1] unfolding oexp_def Let_def s.max2_def mem_Collect_eq by safe metis
next
@@ -685,10 +685,10 @@
show "SUPP f \<noteq> {}"
proof (rule ccontr, unfold not_not)
assume "SUPP f = {}"
- moreover from `f \<in> F` assms(1) have "f \<in> FINFUNC" by blast
+ moreover from \<open>f \<in> F\<close> assms(1) have "f \<in> FINFUNC" by blast
ultimately have "f = const"
by (auto simp: fun_eq_iff support_def FinFunc_def Func_def const_def)
- with assms(2) `f \<in> F` show False by blast
+ with assms(2) \<open>f \<in> F\<close> show False by blast
qed
qed
@@ -787,7 +787,7 @@
and SUPPG: "\<forall>g \<in> G. finite (SUPP g) \<and> SUPP g \<noteq> {} \<and> SUPP g \<subseteq> Field s"
using maxim_isMaxim_support support_not_const by auto
define y' where "y' = s.minim {s.maxim (SUPP f) | f. f \<in> G}"
- from G SUPPG maxG `G \<noteq> {}` have y'min: "s.isMinim {s.maxim (SUPP f) | f. f \<in> G} y'"
+ from G SUPPG maxG \<open>G \<noteq> {}\<close> have y'min: "s.isMinim {s.maxim (SUPP f) | f. f \<in> G} y'"
unfolding y'_def by (intro s.minim_isMinim) (auto simp: s.isMaxim_def)
moreover
have "\<forall>g \<in> G. z \<notin> SUPP g" unfolding support_def G_def by auto
@@ -804,7 +804,7 @@
unfolding s.isMinim_def s.isMaxim_def by auto
with zy have "y' \<noteq> y" "(y', y) \<in> s" using antisymD[OF s.ANTISYM] transD[OF s.TRANS]
by blast+
- moreover from `G \<noteq> {}` have "\<exists>g \<in> G. y' = wo_rel.maxim s (SUPP g)" using y'min
+ moreover from \<open>G \<noteq> {}\<close> have "\<exists>g \<in> G. y' = wo_rel.maxim s (SUPP g)" using y'min
by (auto simp: G_def s.isMinim_def)
ultimately show ?thesis using mp[OF spec[OF mp[OF spec[OF 1]]], of y' G] G by auto
qed simp
@@ -1315,8 +1315,8 @@
fix g h assume gh: "g \<in> FinFunc r s" "h \<in> FinFunc r s" "F g \<noteq> F h"
"let m = s.max_fun_diff g h in (g m, h m) \<in> r"
hence "g \<noteq> h" by auto
- note max_fun_diff_in = rs.max_fun_diff_in[OF `g \<noteq> h` gh(1,2)]
- and max_fun_diff_max = rs.max_fun_diff_max[OF `g \<noteq> h` gh(1,2)]
+ note max_fun_diff_in = rs.max_fun_diff_in[OF \<open>g \<noteq> h\<close> gh(1,2)]
+ and max_fun_diff_max = rs.max_fun_diff_max[OF \<open>g \<noteq> h\<close> gh(1,2)]
with *(4) invff *(2) have "t.max_fun_diff (F g) (F h) = f (s.max_fun_diff g h)"
unfolding t.max_fun_diff_def compat_def
by (intro t.maxim_equality) (auto simp: t.isMaxim_def F_def dest: injfD)
@@ -1401,14 +1401,14 @@
next
case False thus ?thesis
proof (cases "r = {}")
- case True thus ?thesis using t `t \<noteq> {}` st.oexp_Well_order ozero_ordLeq[unfolded ozero_def]
+ case True thus ?thesis using t \<open>t \<noteq> {}\<close> st.oexp_Well_order ozero_ordLeq[unfolded ozero_def]
by auto
next
case False
from assms obtain f where f: "embed r s f" unfolding ordLeq_def by blast
hence f_underS: "\<forall>a\<in>Field r. f a \<in> Field s \<and> f ` underS r a \<subseteq> underS s (f a)"
using embed_in_Field[OF rt.rWELL f] embed_underS2[OF rt.rWELL st.rWELL f] by auto
- from f `t \<noteq> {}` False have *: "Field r \<noteq> {}" "Field s \<noteq> {}" "Field t \<noteq> {}"
+ from f \<open>t \<noteq> {}\<close> False have *: "Field r \<noteq> {}" "Field s \<noteq> {}" "Field t \<noteq> {}"
unfolding Field_def embed_def under_def bij_betw_def by auto
with f obtain x where "s.zero = f x" "x \<in> Field r" unfolding embed_def bij_betw_def
using embed_in_Field[OF r.WELL f] s.zero_under set_mp[OF under_Field[of r]] by blast
@@ -1435,11 +1435,11 @@
by simp metis
moreover
with hg have "t.max_fun_diff (?f h) (?f g) = t.max_fun_diff h g" unfolding rt.oexp_def
- using rt.max_fun_diff[OF `h \<noteq> g`] rt.max_fun_diff_in[OF `h \<noteq> g`]
+ using rt.max_fun_diff[OF \<open>h \<noteq> g\<close>] rt.max_fun_diff_in[OF \<open>h \<noteq> g\<close>]
by (subst t.max_fun_diff_def, intro t.maxim_equality)
(auto simp: t.isMaxim_def intro: inj_onD[OF inj] intro!: rt.max_fun_diff_max)
with Field_fg Field_fh hg fz f_underS compat neq have "(?f h, ?f g) \<in> st.oexp"
- using rt.max_fun_diff[OF `h \<noteq> g`] rt.max_fun_diff_in[OF `h \<noteq> g`] unfolding st.Field_oexp
+ using rt.max_fun_diff[OF \<open>h \<noteq> g\<close>] rt.max_fun_diff_in[OF \<open>h \<noteq> g\<close>] unfolding st.Field_oexp
unfolding rt.oexp_def st.oexp_def Let_def compat_def by auto
ultimately show "?f h \<in> underS (s ^o t) (?f g)" unfolding underS_def by auto
qed
@@ -1497,9 +1497,9 @@
(is "?L = (?R1 \<and> ?R2)")
proof safe
assume ?L
- from `?L` show ?R1 unfolding FinFunc_def Field_osum Func_def Int_iff fin_support_Field_osum o_def
+ from \<open>?L\<close> show ?R1 unfolding FinFunc_def Field_osum Func_def Int_iff fin_support_Field_osum o_def
by (auto split: sum.splits)
- from `?L` show ?R2 unfolding FinFunc_def Field_osum Func_def Int_iff fin_support_Field_osum o_def
+ from \<open>?L\<close> show ?R2 unfolding FinFunc_def Field_osum Func_def Int_iff fin_support_Field_osum o_def
by (auto split: sum.splits)
next
assume ?R1 ?R2
@@ -1674,14 +1674,14 @@
interpret rst: wo_rel2 "r ^o s" t by unfold_locales (rule oexp_Well_order[OF r s], rule t)
show ?thesis
proof (cases "s = {} \<or> t = {}")
- case True with `r = {}` show ?thesis
+ case True with \<open>r = {}\<close> show ?thesis
by (auto simp: oexp_empty[OF oexp_Well_order[OF Well_order_empty s]]
intro!: ordIso_transitive[OF ordIso_symmetric[OF oone_ordIso] oone_ordIso]
ordIso_transitive[OF oone_ordIso_oexp[OF ordIso_symmetric[OF oone_ordIso] t] oone_ordIso])
next
case False
moreover hence "s *o t \<noteq> {}" unfolding oprod_def Field_def by fastforce
- ultimately show ?thesis using `r = {}` ozero_ordIso
+ ultimately show ?thesis using \<open>r = {}\<close> ozero_ordIso
by (auto simp add: s t oprod_Well_order ozero_def)
qed
next
--- a/src/HOL/Cardinals/Wellfounded_More.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Cardinals/Wellfounded_More.thy Thu May 26 17:51:22 2016 +0200
@@ -5,13 +5,13 @@
More on well-founded relations.
*)
-section {* More on Well-Founded Relations *}
+section \<open>More on Well-Founded Relations\<close>
theory Wellfounded_More
imports Wellfounded Order_Relation_More
begin
-subsection {* Well-founded recursion via genuine fixpoints *}
+subsection \<open>Well-founded recursion via genuine fixpoints\<close>
(*2*)lemma adm_wf_unique_fixpoint:
fixes r :: "('a * 'a) set" and
--- a/src/HOL/Cardinals/Wellorder_Constructions.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Cardinals/Wellorder_Constructions.thy Thu May 26 17:51:22 2016 +0200
@@ -5,7 +5,7 @@
Constructions on wellorders.
*)
-section {* Constructions on Wellorders *}
+section \<open>Constructions on Wellorders\<close>
theory Wellorder_Constructions
imports
@@ -23,7 +23,7 @@
lemma Func_emp2[simp]: "A \<noteq> {} \<Longrightarrow> Func A {} = {}" by auto
-subsection {* Restriction to a set *}
+subsection \<open>Restriction to a set\<close>
lemma Restr_incr2:
"r <= r' \<Longrightarrow> Restr r A <= Restr r' A"
@@ -55,7 +55,7 @@
by blast
-subsection {* Order filters versus restrictions and embeddings *}
+subsection \<open>Order filters versus restrictions and embeddings\<close>
lemma ofilter_Restr:
assumes WELL: "Well_order r" and
@@ -92,7 +92,7 @@
by (auto simp add: ofilter_subset_embedS_iso)
-subsection {* Ordering the well-orders by existence of embeddings *}
+subsection \<open>Ordering the well-orders by existence of embeddings\<close>
corollary ordLeq_refl_on: "refl_on {r. Well_order r} ordLeq"
using ordLeq_reflexive unfolding ordLeq_def refl_on_def
@@ -169,7 +169,7 @@
by (auto simp add: ofilter_ordLeq wo_rel.under_ofilter wo_rel_def)
-subsection {* Copy via direct images *}
+subsection \<open>Copy via direct images\<close>
lemma Id_dir_image: "dir_image Id f \<le> Id"
unfolding dir_image_def by auto
@@ -275,9 +275,9 @@
qed
-subsection {* The maxim among a finite set of ordinals *}
+subsection \<open>The maxim among a finite set of ordinals\<close>
-text {* The correct phrasing would be ``a maxim of ...", as @{text "\<le>o"} is only a preorder. *}
+text \<open>The correct phrasing would be ``a maxim of ...", as \<open>\<le>o\<close> is only a preorder.\<close>
definition isOmax :: "'a rel set \<Rightarrow> 'a rel \<Rightarrow> bool"
where
@@ -437,7 +437,7 @@
qed
-subsection {* Limit and succesor ordinals *}
+subsection \<open>Limit and succesor ordinals\<close>
lemma embed_underS2:
assumes r: "Well_order r" and s: "Well_order s" and g: "embed r s g" and a: "a \<in> Field r"
@@ -548,7 +548,7 @@
by (elim cases_Total3 disjE) (auto elim: cases_Total3 intro!: assms simp: Field_def)
qed
-subsubsection {* Successor and limit elements of an ordinal *}
+subsubsection \<open>Successor and limit elements of an ordinal\<close>
definition "succ i \<equiv> suc {i}"
@@ -724,7 +724,7 @@
else L f i"
-subsubsection {* Well-order recursion with (zero), succesor, and limit *}
+subsubsection \<open>Well-order recursion with (zero), succesor, and limit\<close>
definition worecSL :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> (('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b"
where "worecSL S L \<equiv> worec (mergeSL S L)"
@@ -819,7 +819,7 @@
qed
-subsubsection {* Well-order succ-lim induction *}
+subsubsection \<open>Well-order succ-lim induction\<close>
lemma ord_cases:
obtains j where "i = succ j" and "aboveS j \<noteq> {}" | "isLim i"
@@ -901,7 +901,7 @@
abbreviation "worecZSL \<equiv> wo_rel.worecZSL"
-subsection {* Projections of wellorders *}
+subsection \<open>Projections of wellorders\<close>
definition "oproj r s f \<equiv> Field s \<subseteq> f ` (Field r) \<and> compat r s f"
@@ -1018,7 +1018,7 @@
fix b1 assume "b1 \<in> g ` underS r a"
then obtain a1 where a1: "b1 = g a1" and a1: "a1 \<in> underS r a" by auto
hence "b1 \<in> underS s (f a)"
- using a by (metis `\<And>a1. a1 \<in> underS r a \<Longrightarrow> g a1 \<in> underS s (f a)`)
+ using a by (metis \<open>\<And>a1. a1 \<in> underS r a \<Longrightarrow> g a1 \<in> underS s (f a)\<close>)
thus "f a \<noteq> b1 \<and> (b1, f a) \<in> s" unfolding underS_def by auto
qed(insert fa, auto)
thus "g a \<in> under s (f a)" unfolding under_def by auto
@@ -1061,7 +1061,7 @@
by (auto dest!: oproj_Field2[OF f] inv_into_injective intro!: inv_into_into)
ultimately have "(inv_into (Field r) f b, inv_into (Field r) f a) \<in> r"
using r by (auto simp: well_order_on_def linear_order_on_def total_on_def)
- with f[unfolded oproj_def compat_def] *(1) `a \<in> Field s`
+ with f[unfolded oproj_def compat_def] *(1) \<open>a \<in> Field s\<close>
f_inv_into_f[of b f "Field r"] f_inv_into_f[of a f "Field r"]
have "(b, a) \<in> s" by (metis in_mono)
with *(2,3) s have False
--- a/src/HOL/Cardinals/Wellorder_Embedding.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Cardinals/Wellorder_Embedding.thy Thu May 26 17:51:22 2016 +0200
@@ -5,13 +5,13 @@
Well-order embeddings.
*)
-section {* Well-Order Embeddings *}
+section \<open>Well-Order Embeddings\<close>
theory Wellorder_Embedding
imports BNF_Wellorder_Embedding Fun_More Wellorder_Relation
begin
-subsection {* Auxiliaries *}
+subsection \<open>Auxiliaries\<close>
lemma UNION_bij_betw_ofilter:
assumes WELL: "Well_order r" and
@@ -27,8 +27,8 @@
qed
-subsection {* (Well-order) embeddings, strict embeddings, isomorphisms and order-compatible
-functions *}
+subsection \<open>(Well-order) embeddings, strict embeddings, isomorphisms and order-compatible
+functions\<close>
lemma embed_halfcong:
assumes EQ: "\<And> a. a \<in> Field r \<Longrightarrow> f a = g a" and
@@ -103,7 +103,7 @@
using one_set_greater[of UNIV UNIV] by auto
-subsection {* Uniqueness of embeddings *}
+subsection \<open>Uniqueness of embeddings\<close>
lemma comp_embedS:
assumes WELL: "Well_order r" and WELL': "Well_order r'" and WELL'': "Well_order r''"
--- a/src/HOL/Cardinals/Wellorder_Extension.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Cardinals/Wellorder_Extension.thy Thu May 26 17:51:22 2016 +0200
@@ -2,16 +2,16 @@
Author: Christian Sternagel, JAIST
*)
-section {* Extending Well-founded Relations to Wellorders *}
+section \<open>Extending Well-founded Relations to Wellorders\<close>
theory Wellorder_Extension
imports Main Order_Union
begin
-subsection {* Extending Well-founded Relations to Wellorders *}
+subsection \<open>Extending Well-founded Relations to Wellorders\<close>
-text {*A \emph{downset} (also lower set, decreasing set, initial segment, or
-downward closed set) is closed w.r.t.\ smaller elements.*}
+text \<open>A \emph{downset} (also lower set, decreasing set, initial segment, or
+downward closed set) is closed w.r.t.\ smaller elements.\<close>
definition downset_on where
"downset_on A r = (\<forall>x y. (x, y) \<in> r \<and> y \<in> A \<longrightarrow> x \<in> A)"
@@ -30,7 +30,7 @@
"downset_on A r \<Longrightarrow> (x, y) \<in> r \<Longrightarrow> y \<in> A \<Longrightarrow> x \<in> A"
unfolding downset_on_def by blast
-text {*Extensions of relations w.r.t.\ a given set.*}
+text \<open>Extensions of relations w.r.t.\ a given set.\<close>
definition extension_on where
"extension_on A r s = (\<forall>x\<in>A. \<forall>y\<in>A. (x, y) \<in> s \<longrightarrow> (x, y) \<in> r)"
@@ -60,7 +60,7 @@
lemma extension_on_empty [simp]: "extension_on {} p q"
by (auto simp: extension_on_def)
-text {*Every well-founded relation can be extended to a wellorder.*}
+text \<open>Every well-founded relation can be extended to a wellorder.\<close>
theorem well_order_extension:
assumes "wf p"
shows "\<exists>w. p \<subseteq> w \<and> Well_order w"
@@ -79,39 +79,39 @@
trans_def I_def elim: trans_init_seg_of)
{ fix R assume "R \<in> Chains I"
then have Ris: "R \<in> Chains init_seg_of" using mono_Chains [OF I_init] by blast
- have subch: "chain\<^sub>\<subseteq> R" using `R \<in> Chains I` I_init
+ have subch: "chain\<^sub>\<subseteq> R" using \<open>R \<in> Chains I\<close> I_init
by (auto simp: init_seg_of_def chain_subset_def Chains_def)
have "\<forall>r\<in>R. Refl r" and "\<forall>r\<in>R. trans r" and "\<forall>r\<in>R. antisym r" and
"\<forall>r\<in>R. Total r" and "\<forall>r\<in>R. wf (r - Id)" and
"\<And>r. r \<in> R \<Longrightarrow> downset_on (Field r) p" and
"\<And>r. r \<in> R \<Longrightarrow> extension_on (Field r) r p"
- using Chains_wo [OF `R \<in> Chains I`] by (simp_all add: order_on_defs)
- have "Refl (\<Union>R)" using `\<forall>r\<in>R. Refl r` unfolding refl_on_def by fastforce
+ using Chains_wo [OF \<open>R \<in> Chains I\<close>] by (simp_all add: order_on_defs)
+ have "Refl (\<Union>R)" using \<open>\<forall>r\<in>R. Refl r\<close> unfolding refl_on_def by fastforce
moreover have "trans (\<Union>R)"
- by (rule chain_subset_trans_Union [OF subch `\<forall>r\<in>R. trans r`])
+ by (rule chain_subset_trans_Union [OF subch \<open>\<forall>r\<in>R. trans r\<close>])
moreover have "antisym (\<Union>R)"
- by (rule chain_subset_antisym_Union [OF subch `\<forall>r\<in>R. antisym r`])
+ by (rule chain_subset_antisym_Union [OF subch \<open>\<forall>r\<in>R. antisym r\<close>])
moreover have "Total (\<Union>R)"
- by (rule chain_subset_Total_Union [OF subch `\<forall>r\<in>R. Total r`])
+ by (rule chain_subset_Total_Union [OF subch \<open>\<forall>r\<in>R. Total r\<close>])
moreover have "wf ((\<Union>R) - Id)"
proof -
have "(\<Union>R) - Id = \<Union>{r - Id | r. r \<in> R}" by blast
- with `\<forall>r\<in>R. wf (r - Id)` wf_Union_wf_init_segs [OF Chains_inits_DiffI [OF Ris]]
+ with \<open>\<forall>r\<in>R. wf (r - Id)\<close> wf_Union_wf_init_segs [OF Chains_inits_DiffI [OF Ris]]
show ?thesis by fastforce
qed
ultimately have "Well_order (\<Union>R)" by (simp add: order_on_defs)
moreover have "\<forall>r\<in>R. r initial_segment_of \<Union>R" using Ris
by (simp add: Chains_init_seg_of_Union)
moreover have "downset_on (Field (\<Union>R)) p"
- by (rule downset_on_Union [OF `\<And>r. r \<in> R \<Longrightarrow> downset_on (Field r) p`])
+ by (rule downset_on_Union [OF \<open>\<And>r. r \<in> R \<Longrightarrow> downset_on (Field r) p\<close>])
moreover have "extension_on (Field (\<Union>R)) (\<Union>R) p"
- by (rule chain_subset_extension_on_Union [OF subch `\<And>r. r \<in> R \<Longrightarrow> extension_on (Field r) r p`])
+ by (rule chain_subset_extension_on_Union [OF subch \<open>\<And>r. r \<in> R \<Longrightarrow> extension_on (Field r) r p\<close>])
ultimately have "\<Union>R \<in> ?K \<and> (\<forall>r\<in>R. (r,\<Union>R) \<in> I)"
- using mono_Chains [OF I_init] and `R \<in> Chains I`
+ using mono_Chains [OF I_init] and \<open>R \<in> Chains I\<close>
by (simp (no_asm) add: I_def del: Field_Union) (metis Chains_wo)
}
then have 1: "\<forall>R\<in>Chains I. \<exists>u\<in>Field I. \<forall>r\<in>R. (r, u) \<in> I" by (subst FI) blast
- txt {*Zorn's Lemma yields a maximal wellorder m.*}
+ txt \<open>Zorn's Lemma yields a maximal wellorder m.\<close>
from Zorns_po_lemma [OF 0 1] obtain m :: "('a \<times> 'a) set"
where "Well_order m" and "downset_on (Field m) p" and "extension_on (Field m) m p" and
max: "\<forall>r. Well_order r \<and> downset_on (Field r) p \<and> extension_on (Field r) r p \<and>
@@ -124,52 +124,52 @@
with assms [unfolded wf_eq_minimal, THEN spec, of ?Q]
obtain x where "x \<in> Field p" and "x \<notin> Field m" and
min: "\<forall>y. (y, x) \<in> p \<longrightarrow> y \<notin> ?Q" by blast
- txt {*Add @{term x} as topmost element to @{term m}.*}
+ txt \<open>Add @{term x} as topmost element to @{term m}.\<close>
let ?s = "{(y, x) | y. y \<in> Field m}"
let ?m = "insert (x, x) m \<union> ?s"
have Fm: "Field ?m = insert x (Field m)" by (auto simp: Field_def)
have "Refl m" and "trans m" and "antisym m" and "Total m" and "wf (m - Id)"
- using `Well_order m` by (simp_all add: order_on_defs)
- txt {*We show that the extension is a wellorder.*}
- have "Refl ?m" using `Refl m` Fm by (auto simp: refl_on_def)
- moreover have "trans ?m" using `trans m` `x \<notin> Field m`
+ using \<open>Well_order m\<close> by (simp_all add: order_on_defs)
+ txt \<open>We show that the extension is a wellorder.\<close>
+ have "Refl ?m" using \<open>Refl m\<close> Fm by (auto simp: refl_on_def)
+ moreover have "trans ?m" using \<open>trans m\<close> \<open>x \<notin> Field m\<close>
unfolding trans_def Field_def Domain_unfold Domain_converse [symmetric] by blast
- moreover have "antisym ?m" using `antisym m` `x \<notin> Field m`
+ moreover have "antisym ?m" using \<open>antisym m\<close> \<open>x \<notin> Field m\<close>
unfolding antisym_def Field_def Domain_unfold Domain_converse [symmetric] by blast
- moreover have "Total ?m" using `Total m` Fm by (auto simp: Relation.total_on_def)
+ moreover have "Total ?m" using \<open>Total m\<close> Fm by (auto simp: Relation.total_on_def)
moreover have "wf (?m - Id)"
proof -
- have "wf ?s" using `x \<notin> Field m`
+ have "wf ?s" using \<open>x \<notin> Field m\<close>
by (simp add: wf_eq_minimal Field_def Domain_unfold Domain_converse [symmetric]) metis
- thus ?thesis using `wf (m - Id)` `x \<notin> Field m`
- wf_subset [OF `wf ?s` Diff_subset]
+ thus ?thesis using \<open>wf (m - Id)\<close> \<open>x \<notin> Field m\<close>
+ wf_subset [OF \<open>wf ?s\<close> Diff_subset]
by (fastforce intro!: wf_Un simp add: Un_Diff Field_def)
qed
ultimately have "Well_order ?m" by (simp add: order_on_defs)
moreover have "extension_on (Field ?m) ?m p"
- using `extension_on (Field m) m p` `downset_on (Field m) p`
+ using \<open>extension_on (Field m) m p\<close> \<open>downset_on (Field m) p\<close>
by (subst Fm) (auto simp: extension_on_def dest: downset_onD)
moreover have "downset_on (Field ?m) p"
apply (subst Fm)
- using `downset_on (Field m) p` and min
+ using \<open>downset_on (Field m) p\<close> and min
unfolding downset_on_def Field_def by blast
moreover have "(m, ?m) \<in> I"
- using `Well_order m` and `Well_order ?m` and
- `downset_on (Field m) p` and `downset_on (Field ?m) p` and
- `extension_on (Field m) m p` and `extension_on (Field ?m) ?m p` and
- `Refl m` and `x \<notin> Field m`
+ using \<open>Well_order m\<close> and \<open>Well_order ?m\<close> and
+ \<open>downset_on (Field m) p\<close> and \<open>downset_on (Field ?m) p\<close> and
+ \<open>extension_on (Field m) m p\<close> and \<open>extension_on (Field ?m) ?m p\<close> and
+ \<open>Refl m\<close> and \<open>x \<notin> Field m\<close>
by (auto simp: I_def init_seg_of_def refl_on_def)
ultimately
- --{*This contradicts maximality of m:*}
- show False using max and `x \<notin> Field m` unfolding Field_def by blast
+ \<comment>\<open>This contradicts maximality of m:\<close>
+ show False using max and \<open>x \<notin> Field m\<close> unfolding Field_def by blast
qed
have "p \<subseteq> m"
- using `Field p \<subseteq> Field m` and `extension_on (Field m) m p`
+ using \<open>Field p \<subseteq> Field m\<close> and \<open>extension_on (Field m) m p\<close>
unfolding Field_def extension_on_def by auto fast
- with `Well_order m` show ?thesis by blast
+ with \<open>Well_order m\<close> show ?thesis by blast
qed
-text {*Every well-founded relation can be extended to a total wellorder.*}
+text \<open>Every well-founded relation can be extended to a total wellorder.\<close>
corollary total_well_order_extension:
assumes "wf p"
shows "\<exists>w. p \<subseteq> w \<and> Well_order w \<and> Field w = UNIV"
@@ -181,7 +181,7 @@
have [simp]: "Field w' = ?A" using well_order_on_Well_order [OF wo'] by simp
have *: "Field w \<inter> Field w' = {}" by simp
let ?w = "w \<union>o w'"
- have "p \<subseteq> ?w" using `p \<subseteq> w` by (auto simp: Osum_def)
+ have "p \<subseteq> ?w" using \<open>p \<subseteq> w\<close> by (auto simp: Osum_def)
moreover have "Well_order ?w" using Osum_Well_order [OF * wo] and wo' by simp
moreover have "Field ?w = UNIV" by (simp add: Field_Osum)
ultimately show ?thesis by blast
@@ -191,23 +191,23 @@
assumes "wf p" and "Field p \<subseteq> A"
shows "\<exists>w. p \<subseteq> w \<and> well_order_on A w"
proof -
- from total_well_order_extension [OF `wf p`] obtain r
+ from total_well_order_extension [OF \<open>wf p\<close>] obtain r
where "p \<subseteq> r" and wo: "Well_order r" and univ: "Field r = UNIV" by blast
let ?r = "{(x, y). x \<in> A \<and> y \<in> A \<and> (x, y) \<in> r}"
- from `p \<subseteq> r` have "p \<subseteq> ?r" using `Field p \<subseteq> A` by (auto simp: Field_def)
+ from \<open>p \<subseteq> r\<close> have "p \<subseteq> ?r" using \<open>Field p \<subseteq> A\<close> by (auto simp: Field_def)
have 1: "Field ?r = A" using wo univ
by (fastforce simp: Field_def order_on_defs refl_on_def)
have "Refl r" "trans r" "antisym r" "Total r" "wf (r - Id)"
- using `Well_order r` by (simp_all add: order_on_defs)
- have "refl_on A ?r" using `Refl r` by (auto simp: refl_on_def univ)
- moreover have "trans ?r" using `trans r`
+ using \<open>Well_order r\<close> by (simp_all add: order_on_defs)
+ have "refl_on A ?r" using \<open>Refl r\<close> by (auto simp: refl_on_def univ)
+ moreover have "trans ?r" using \<open>trans r\<close>
unfolding trans_def by blast
- moreover have "antisym ?r" using `antisym r`
+ moreover have "antisym ?r" using \<open>antisym r\<close>
unfolding antisym_def by blast
- moreover have "total_on A ?r" using `Total r` by (simp add: total_on_def univ)
- moreover have "wf (?r - Id)" by (rule wf_subset [OF `wf(r - Id)`]) blast
+ moreover have "total_on A ?r" using \<open>Total r\<close> by (simp add: total_on_def univ)
+ moreover have "wf (?r - Id)" by (rule wf_subset [OF \<open>wf(r - Id)\<close>]) blast
ultimately have "well_order_on A ?r" by (simp add: order_on_defs)
- with `p \<subseteq> ?r` show ?thesis by blast
+ with \<open>p \<subseteq> ?r\<close> show ?thesis by blast
qed
end
--- a/src/HOL/Cardinals/Wellorder_Relation.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Cardinals/Wellorder_Relation.thy Thu May 26 17:51:22 2016 +0200
@@ -5,7 +5,7 @@
Well-order relations.
*)
-section {* Well-Order Relations *}
+section \<open>Well-Order Relations\<close>
theory Wellorder_Relation
imports BNF_Wellorder_Relation Wellfounded_More
@@ -14,7 +14,7 @@
context wo_rel
begin
-subsection {* Auxiliaries *}
+subsection \<open>Auxiliaries\<close>
lemma PREORD: "Preorder r"
using WELL order_on_defs[of _ r] by auto
@@ -29,7 +29,7 @@
using TOTALS by auto
-subsection {* Well-founded induction and recursion adapted to non-strict well-order relations *}
+subsection \<open>Well-founded induction and recursion adapted to non-strict well-order relations\<close>
lemma worec_unique_fixpoint:
assumes ADM: "adm_wo H" and fp: "f = H f"
@@ -44,7 +44,7 @@
qed
-subsubsection {* Properties of max2 *}
+subsubsection \<open>Properties of max2\<close>
lemma max2_iff:
assumes "a \<in> Field r" and "b \<in> Field r"
@@ -60,7 +60,7 @@
qed
-subsubsection {* Properties of minim *}
+subsubsection \<open>Properties of minim\<close>
lemma minim_Under:
"\<lbrakk>B \<le> Field r; B \<noteq> {}\<rbrakk> \<Longrightarrow> minim B \<in> Under B"
@@ -166,7 +166,7 @@
qed
-subsubsection {* Properties of supr *}
+subsubsection \<open>Properties of supr\<close>
lemma supr_Above:
assumes SUB: "B \<le> Field r" and ABOVE: "Above B \<noteq> {}"
@@ -287,7 +287,7 @@
qed
-subsubsection {* Properties of successor *}
+subsubsection \<open>Properties of successor\<close>
lemma suc_least:
"\<lbrakk>B \<le> Field r; a \<in> Field r; (\<And> b. b \<in> B \<Longrightarrow> a \<noteq> b \<and> (b,a) \<in> r)\<rbrakk>
@@ -400,7 +400,7 @@
qed
-subsubsection {* Properties of order filters *}
+subsubsection \<open>Properties of order filters\<close>
lemma ofilter_Under[simp]:
assumes "A \<le> Field r"
@@ -453,7 +453,7 @@
using ofilter_under_UNION [of A] by auto
-subsubsection {* Other properties *}
+subsubsection \<open>Other properties\<close>
lemma Trans_Under_regressive:
assumes NE: "A \<noteq> {}" and SUB: "A \<le> Field r"
--- a/src/HOL/Codegenerator_Test/Candidates.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Codegenerator_Test/Candidates.thy Thu May 26 17:51:22 2016 +0200
@@ -1,7 +1,7 @@
(* Author: Florian Haftmann, TU Muenchen *)
-section {* A huge collection of equations to generate code from *}
+section \<open>A huge collection of equations to generate code from\<close>
theory Candidates
imports
@@ -40,7 +40,7 @@
code_reserved SML upto
-text \<open>Explicit check in @{text OCaml} for correct precedence of let expressions in list expressions\<close>
+text \<open>Explicit check in \<open>OCaml\<close> for correct precedence of let expressions in list expressions\<close>
definition funny_list :: "bool list"
where
@@ -58,7 +58,7 @@
where
"check_list = (if funny_list = funny_list' then () else undefined)"
-text \<open>Explicit check in @{text Scala} for correct bracketing of abstractions\<close>
+text \<open>Explicit check in \<open>Scala\<close> for correct bracketing of abstractions\<close>
definition funny_funs :: "(bool \<Rightarrow> bool) list \<Rightarrow> (bool \<Rightarrow> bool) list"
where
--- a/src/HOL/Codegenerator_Test/Code_Test_GHC.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Codegenerator_Test/Code_Test_GHC.thy Thu May 26 17:51:22 2016 +0200
@@ -10,7 +10,7 @@
value [GHC] "14 + 7 * -12 :: integer"
-test_code -- \<open>Tests for the serialisation of @{const gcd} on @{typ integer}\<close>
+test_code \<comment> \<open>Tests for the serialisation of @{const gcd} on @{typ integer}\<close>
"gcd 15 10 = (5 :: integer)"
"gcd 15 (- 10) = (5 :: integer)"
"gcd (- 10) 15 = (5 :: integer)"
--- a/src/HOL/Codegenerator_Test/Code_Test_OCaml.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Codegenerator_Test/Code_Test_OCaml.thy Thu May 26 17:51:22 2016 +0200
@@ -10,7 +10,7 @@
value [OCaml] "14 + 7 * -12 :: integer"
-test_code -- \<open>Tests for the serialisation of @{const gcd} on @{typ integer}\<close>
+test_code \<comment> \<open>Tests for the serialisation of @{const gcd} on @{typ integer}\<close>
"gcd 15 10 = (5 :: integer)"
"gcd 15 (- 10) = (5 :: integer)"
"gcd (- 10) 15 = (5 :: integer)"
--- a/src/HOL/Codegenerator_Test/Code_Test_Scala.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Codegenerator_Test/Code_Test_Scala.thy Thu May 26 17:51:22 2016 +0200
@@ -12,7 +12,7 @@
value [Scala] "14 + 7 * -12 :: integer"
-test_code -- \<open>Tests for the serialisation of @{const gcd} on @{typ integer}\<close>
+test_code \<comment> \<open>Tests for the serialisation of @{const gcd} on @{typ integer}\<close>
"gcd 15 10 = (5 :: integer)"
"gcd 15 (- 10) = (5 :: integer)"
"gcd (- 10) 15 = (5 :: integer)"
--- a/src/HOL/Codegenerator_Test/Generate.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Codegenerator_Test/Generate.thy Thu May 26 17:51:22 2016 +0200
@@ -1,7 +1,7 @@
(* Author: Florian Haftmann, TU Muenchen *)
-section {* Pervasive test of code generator *}
+section \<open>Pervasive test of code generator\<close>
theory Generate
imports
@@ -10,10 +10,10 @@
"~~/src/HOL/Library/Finite_Lattice"
begin
-text {*
+text \<open>
If any of the checks fails, inspect the code generated
- by a corresponding @{text export_code} command.
-*}
+ by a corresponding \<open>export_code\<close> command.
+\<close>
export_code _ checking SML OCaml? Haskell? Scala
--- a/src/HOL/Codegenerator_Test/Generate_Binary_Nat.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Codegenerator_Test/Generate_Binary_Nat.thy Thu May 26 17:51:22 2016 +0200
@@ -1,7 +1,7 @@
(* Author: Florian Haftmann, TU Muenchen *)
-section {* Pervasive test of code generator *}
+section \<open>Pervasive test of code generator\<close>
theory Generate_Binary_Nat
imports
@@ -11,10 +11,10 @@
"~~/src/HOL/Library/Code_Binary_Nat"
begin
-text {*
+text \<open>
If any of the checks fails, inspect the code generated
- by a corresponding @{text export_code} command.
-*}
+ by a corresponding \<open>export_code\<close> command.
+\<close>
export_code _ checking SML OCaml? Haskell? Scala
--- a/src/HOL/Codegenerator_Test/Generate_Efficient_Datastructures.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Codegenerator_Test/Generate_Efficient_Datastructures.thy Thu May 26 17:51:22 2016 +0200
@@ -1,7 +1,7 @@
(* Author: Ondrej Kuncar, TU Muenchen *)
-section {* Pervasive test of code generator *}
+section \<open>Pervasive test of code generator\<close>
theory Generate_Efficient_Datastructures
imports
@@ -18,7 +18,7 @@
val consts = map_filter (try (curry (Axclass.param_of_inst thy)
@{const_name "Quickcheck_Narrowing.partial_term_of"})) tycos;
in fold Code.del_eqns consts thy end
-\<close> -- \<open>drop technical stuff from @{text Quickcheck_Narrowing} which is tailored towards Haskell\<close>
+\<close> \<comment> \<open>drop technical stuff from \<open>Quickcheck_Narrowing\<close> which is tailored towards Haskell\<close>
(*
The following code equations have to be deleted because they use
--- a/src/HOL/Codegenerator_Test/Generate_Pretty_Char.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Codegenerator_Test/Generate_Pretty_Char.thy Thu May 26 17:51:22 2016 +0200
@@ -1,7 +1,7 @@
(* Author: Florian Haftmann, TU Muenchen *)
-section {* Pervasive test of code generator *}
+section \<open>Pervasive test of code generator\<close>
theory Generate_Pretty_Char
imports
@@ -11,10 +11,10 @@
"~~/src/HOL/Library/Code_Char"
begin
-text {*
+text \<open>
If any of the checks fails, inspect the code generated
- by a corresponding @{text export_code} command.
-*}
+ by a corresponding \<open>export_code\<close> command.
+\<close>
export_code _ checking SML OCaml? Haskell? Scala
--- a/src/HOL/Codegenerator_Test/Generate_Target_Nat.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Codegenerator_Test/Generate_Target_Nat.thy Thu May 26 17:51:22 2016 +0200
@@ -1,7 +1,7 @@
(* Author: Florian Haftmann, TU Muenchen *)
-section {* Pervasive test of code generator *}
+section \<open>Pervasive test of code generator\<close>
theory Generate_Target_Nat
imports
@@ -11,10 +11,10 @@
"~~/src/HOL/Library/Code_Target_Numeral"
begin
-text {*
+text \<open>
If any of the checks fails, inspect the code generated
- by a corresponding @{text export_code} command.
-*}
+ by a corresponding \<open>export_code\<close> command.
+\<close>
export_code _ checking SML OCaml? Haskell? Scala
--- a/src/HOL/Corec_Examples/LFilter.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Corec_Examples/LFilter.thy Thu May 26 17:51:22 2016 +0200
@@ -82,7 +82,7 @@
else if P (lhd xs) then LCons (lhd xs) (f (ltl xs))
else f (ltl xs))"
shows "f = lfilter P"
--- \<open>It seems as if we cannot use @{thm lfilter_unique_weak} for showing this as the induction and the coinduction must be nested\<close>
+\<comment> \<open>It seems as if we cannot use @{thm lfilter_unique_weak} for showing this as the induction and the coinduction must be nested\<close>
proof(rule ext)
show "f xs = lfilter P xs" for xs
proof(coinduction arbitrary: xs)
--- a/src/HOL/Datatype_Examples/Compat.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Datatype_Examples/Compat.thy Thu May 26 17:51:22 2016 +0200
@@ -5,15 +5,15 @@
Tests for compatibility with the old datatype package.
*)
-section {* Tests for Compatibility with the Old Datatype Package *}
+section \<open>Tests for Compatibility with the Old Datatype Package\<close>
theory Compat
imports "~~/src/HOL/Library/Old_Datatype"
begin
-subsection {* Viewing and Registering New-Style Datatypes as Old-Style Ones *}
+subsection \<open>Viewing and Registering New-Style Datatypes as Old-Style Ones\<close>
-ML {*
+ML \<open>
fun check_len n xs label =
length xs = n orelse error ("Expected length " ^ string_of_int (length xs) ^ " for " ^ label);
@@ -25,13 +25,13 @@
these (Option.map #descr (BNF_LFP_Compat.get_info thy [] T_name)),
these (Option.map #descr (BNF_LFP_Compat.get_info thy [BNF_LFP_Compat.Keep_Nesting] T_name)))
|> tap (check_lens lens);
-*}
+\<close>
old_datatype 'a old_lst = Old_Nl | Old_Cns 'a "'a old_lst"
-text {*
-A few tests to make sure that @{text old_datatype} works as expected:
-*}
+text \<open>
+A few tests to make sure that \<open>old_datatype\<close> works as expected:
+\<close>
primrec old_len :: "'a old_lst \<Rightarrow> nat" where
"old_len Old_Nl = 0"
@@ -49,208 +49,208 @@
lemma "old_len xs = size xs"
by (induct xs) auto
-ML {* get_descrs @{theory} (1, 1, 1) @{type_name old_lst} *}
+ML \<open>get_descrs @{theory} (1, 1, 1) @{type_name old_lst}\<close>
datatype 'a lst = Nl | Cns 'a "'a lst"
-ML {* get_descrs @{theory} (0, 1, 1) @{type_name lst} *}
+ML \<open>get_descrs @{theory} (0, 1, 1) @{type_name lst}\<close>
datatype_compat lst
-ML {* get_descrs @{theory} (1, 1, 1) @{type_name lst} *}
+ML \<open>get_descrs @{theory} (1, 1, 1) @{type_name lst}\<close>
datatype 'b w = W | W' "'b w \<times> 'b list"
-ML {* get_descrs @{theory} (0, 1, 1) @{type_name w} *}
+ML \<open>get_descrs @{theory} (0, 1, 1) @{type_name w}\<close>
datatype_compat w
-ML {* get_descrs @{theory} (2, 2, 1) @{type_name w} *}
+ML \<open>get_descrs @{theory} (2, 2, 1) @{type_name w}\<close>
datatype ('c, 'b) s = L 'c | R 'b
-ML {* get_descrs @{theory} (0, 1, 1) @{type_name s} *}
+ML \<open>get_descrs @{theory} (0, 1, 1) @{type_name s}\<close>
datatype 'd x = X | X' "('d x lst, 'd list) s"
-ML {* get_descrs @{theory} (0, 1, 1) @{type_name x} *}
+ML \<open>get_descrs @{theory} (0, 1, 1) @{type_name x}\<close>
datatype_compat s
-ML {* get_descrs @{theory} (1, 1, 1) @{type_name s} *}
-ML {* get_descrs @{theory} (0, 1, 1) @{type_name x} *}
+ML \<open>get_descrs @{theory} (1, 1, 1) @{type_name s}\<close>
+ML \<open>get_descrs @{theory} (0, 1, 1) @{type_name x}\<close>
datatype_compat x
-ML {* get_descrs @{theory} (3, 3, 1) @{type_name x} *}
+ML \<open>get_descrs @{theory} (3, 3, 1) @{type_name x}\<close>
thm x.induct x.rec
thm compat_x.induct compat_x.rec
datatype 'a tttre = TTTre 'a "'a tttre lst lst lst"
-ML {* get_descrs @{theory} (0, 1, 1) @{type_name tttre} *}
+ML \<open>get_descrs @{theory} (0, 1, 1) @{type_name tttre}\<close>
datatype_compat tttre
-ML {* get_descrs @{theory} (4, 4, 1) @{type_name tttre} *}
+ML \<open>get_descrs @{theory} (4, 4, 1) @{type_name tttre}\<close>
thm tttre.induct tttre.rec
thm compat_tttre.induct compat_tttre.rec
datatype 'a ftre = FEmp | FTre "'a \<Rightarrow> 'a ftre lst"
-ML {* get_descrs @{theory} (0, 1, 1) @{type_name ftre} *}
+ML \<open>get_descrs @{theory} (0, 1, 1) @{type_name ftre}\<close>
datatype_compat ftre
-ML {* get_descrs @{theory} (2, 2, 1) @{type_name ftre} *}
+ML \<open>get_descrs @{theory} (2, 2, 1) @{type_name ftre}\<close>
thm ftre.induct ftre.rec
thm compat_ftre.induct compat_ftre.rec
datatype 'a btre = BTre 'a "'a btre lst" "'a btre lst"
-ML {* get_descrs @{theory} (0, 1, 1) @{type_name btre} *}
+ML \<open>get_descrs @{theory} (0, 1, 1) @{type_name btre}\<close>
datatype_compat btre
-ML {* get_descrs @{theory} (3, 3, 1) @{type_name btre} *}
+ML \<open>get_descrs @{theory} (3, 3, 1) @{type_name btre}\<close>
thm btre.induct btre.rec
thm compat_btre.induct compat_btre.rec
datatype 'a foo = Foo | Foo' 'a "'a bar" and 'a bar = Bar | Bar' 'a "'a foo"
-ML {* get_descrs @{theory} (0, 2, 2) @{type_name foo} *}
-ML {* get_descrs @{theory} (0, 2, 2) @{type_name bar} *}
+ML \<open>get_descrs @{theory} (0, 2, 2) @{type_name foo}\<close>
+ML \<open>get_descrs @{theory} (0, 2, 2) @{type_name bar}\<close>
datatype_compat foo bar
-ML {* get_descrs @{theory} (2, 2, 2) @{type_name foo} *}
-ML {* get_descrs @{theory} (2, 2, 2) @{type_name bar} *}
+ML \<open>get_descrs @{theory} (2, 2, 2) @{type_name foo}\<close>
+ML \<open>get_descrs @{theory} (2, 2, 2) @{type_name bar}\<close>
datatype 'a tre = Tre 'a "'a tre lst"
-ML {* get_descrs @{theory} (0, 1, 1) @{type_name tre} *}
+ML \<open>get_descrs @{theory} (0, 1, 1) @{type_name tre}\<close>
datatype_compat tre
-ML {* get_descrs @{theory} (2, 2, 1) @{type_name tre} *}
+ML \<open>get_descrs @{theory} (2, 2, 1) @{type_name tre}\<close>
thm tre.induct tre.rec
thm compat_tre.induct compat_tre.rec
datatype 'a f = F 'a and 'a g = G 'a
-ML {* get_descrs @{theory} (0, 2, 2) @{type_name f} *}
-ML {* get_descrs @{theory} (0, 2, 2) @{type_name g} *}
+ML \<open>get_descrs @{theory} (0, 2, 2) @{type_name f}\<close>
+ML \<open>get_descrs @{theory} (0, 2, 2) @{type_name g}\<close>
datatype h = H "h f" | H'
-ML {* get_descrs @{theory} (0, 1, 1) @{type_name h} *}
+ML \<open>get_descrs @{theory} (0, 1, 1) @{type_name h}\<close>
datatype_compat f g
-ML {* get_descrs @{theory} (2, 2, 2) @{type_name f} *}
-ML {* get_descrs @{theory} (2, 2, 2) @{type_name g} *}
-ML {* get_descrs @{theory} (0, 1, 1) @{type_name h} *}
+ML \<open>get_descrs @{theory} (2, 2, 2) @{type_name f}\<close>
+ML \<open>get_descrs @{theory} (2, 2, 2) @{type_name g}\<close>
+ML \<open>get_descrs @{theory} (0, 1, 1) @{type_name h}\<close>
datatype_compat h
-ML {* get_descrs @{theory} (3, 3, 1) @{type_name h} *}
+ML \<open>get_descrs @{theory} (3, 3, 1) @{type_name h}\<close>
thm h.induct h.rec
thm compat_h.induct compat_h.rec
datatype myunit = MyUnity
-ML {* get_descrs @{theory} (0, 1, 1) @{type_name myunit} *}
+ML \<open>get_descrs @{theory} (0, 1, 1) @{type_name myunit}\<close>
datatype_compat myunit
-ML {* get_descrs @{theory} (1, 1, 1) @{type_name myunit} *}
+ML \<open>get_descrs @{theory} (1, 1, 1) @{type_name myunit}\<close>
datatype mylist = MyNil | MyCons nat mylist
-ML {* get_descrs @{theory} (0, 1, 1) @{type_name mylist} *}
+ML \<open>get_descrs @{theory} (0, 1, 1) @{type_name mylist}\<close>
datatype_compat mylist
-ML {* get_descrs @{theory} (1, 1, 1) @{type_name mylist} *}
+ML \<open>get_descrs @{theory} (1, 1, 1) @{type_name mylist}\<close>
datatype foo' = FooNil | FooCons bar' foo' and bar' = Bar
-ML {* get_descrs @{theory} (0, 2, 2) @{type_name foo'} *}
-ML {* get_descrs @{theory} (0, 2, 2) @{type_name bar'} *}
+ML \<open>get_descrs @{theory} (0, 2, 2) @{type_name foo'}\<close>
+ML \<open>get_descrs @{theory} (0, 2, 2) @{type_name bar'}\<close>
datatype_compat bar' foo'
-ML {* get_descrs @{theory} (2, 2, 2) @{type_name foo'} *}
-ML {* get_descrs @{theory} (2, 2, 2) @{type_name bar'} *}
+ML \<open>get_descrs @{theory} (2, 2, 2) @{type_name foo'}\<close>
+ML \<open>get_descrs @{theory} (2, 2, 2) @{type_name bar'}\<close>
old_datatype funky = Funky "funky tre" | Funky'
-ML {* get_descrs @{theory} (3, 3, 3) @{type_name funky} *}
+ML \<open>get_descrs @{theory} (3, 3, 3) @{type_name funky}\<close>
old_datatype fnky = Fnky "nat tre"
-ML {* get_descrs @{theory} (1, 1, 1) @{type_name fnky} *}
+ML \<open>get_descrs @{theory} (1, 1, 1) @{type_name fnky}\<close>
datatype tree = Tree "tree foo"
-ML {* get_descrs @{theory} (0, 1, 1) @{type_name tree} *}
+ML \<open>get_descrs @{theory} (0, 1, 1) @{type_name tree}\<close>
datatype_compat tree
-ML {* get_descrs @{theory} (3, 3, 1) @{type_name tree} *}
+ML \<open>get_descrs @{theory} (3, 3, 1) @{type_name tree}\<close>
thm tree.induct tree.rec
thm compat_tree.induct compat_tree.rec
-subsection {* Creating New-Style Datatypes Using Old-Style Interfaces *}
+subsection \<open>Creating New-Style Datatypes Using Old-Style Interfaces\<close>
-ML {*
+ML \<open>
val l_specs =
[((@{binding l}, [("'a", @{sort type})], NoSyn),
[(@{binding N}, [], NoSyn),
(@{binding C}, [@{typ 'a}, Type (Sign.full_name @{theory} @{binding l}, [@{typ 'a}])],
NoSyn)])];
-*}
+\<close>
-setup {* snd o BNF_LFP_Compat.add_datatype [] l_specs *}
+setup \<open>snd o BNF_LFP_Compat.add_datatype [] l_specs\<close>
-ML {* get_descrs @{theory} (1, 1, 1) @{type_name l} *}
+ML \<open>get_descrs @{theory} (1, 1, 1) @{type_name l}\<close>
thm l.exhaust l.map l.induct l.rec l.size
-ML {*
+ML \<open>
val t_specs =
[((@{binding t}, [("'b", @{sort type})], NoSyn),
[(@{binding T}, [@{typ 'b},
Type (@{type_name l}, [Type (Sign.full_name @{theory} @{binding t}, [@{typ 'b}])])],
NoSyn)])];
-*}
+\<close>
-setup {* snd o BNF_LFP_Compat.add_datatype [] t_specs *}
+setup \<open>snd o BNF_LFP_Compat.add_datatype [] t_specs\<close>
-ML {* get_descrs @{theory} (2, 2, 1) @{type_name t} *}
+ML \<open>get_descrs @{theory} (2, 2, 1) @{type_name t}\<close>
thm t.exhaust t.map t.induct t.rec t.size
thm compat_t.induct compat_t.rec
-ML {*
+ML \<open>
val ft_specs =
[((@{binding ft}, [("'a", @{sort type})], NoSyn),
[(@{binding FT0}, [], NoSyn),
(@{binding FT}, [@{typ 'a} --> Type (Sign.full_name @{theory} @{binding ft}, [@{typ 'a}])],
NoSyn)])];
-*}
+\<close>
-setup {* snd o BNF_LFP_Compat.add_datatype [] ft_specs *}
+setup \<open>snd o BNF_LFP_Compat.add_datatype [] ft_specs\<close>
-ML {* get_descrs @{theory} (1, 1, 1) @{type_name ft} *}
+ML \<open>get_descrs @{theory} (1, 1, 1) @{type_name ft}\<close>
thm ft.exhaust ft.induct ft.rec ft.size
thm compat_ft.induct compat_ft.rec
--- a/src/HOL/Datatype_Examples/Derivation_Trees/DTree.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Datatype_Examples/Derivation_Trees/DTree.thy Thu May 26 17:51:22 2016 +0200
@@ -5,7 +5,7 @@
Derivation trees with nonterminal internal nodes and terminal leaves.
*)
-section {* Trees with Nonterminal Internal Nodes and Terminal Leaves *}
+section \<open>Trees with Nonterminal Internal Nodes and Terminal Leaves\<close>
theory DTree
imports Prelim
@@ -16,7 +16,7 @@
codatatype dtree = NNode (root: N) (ccont: "(T + dtree) fset")
-subsection{* Transporting the Characteristic Lemmas from @{text "fset"} to @{text "set"} *}
+subsection\<open>Transporting the Characteristic Lemmas from \<open>fset\<close> to \<open>set\<close>\<close>
definition "Node n as \<equiv> NNode n (the_inv fset as)"
definition "cont \<equiv> fset o ccont"
--- a/src/HOL/Datatype_Examples/Derivation_Trees/Gram_Lang.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Datatype_Examples/Derivation_Trees/Gram_Lang.thy Thu May 26 17:51:22 2016 +0200
@@ -5,7 +5,7 @@
Language of a grammar.
*)
-section {* Language of a Grammar *}
+section \<open>Language of a Grammar\<close>
theory Gram_Lang
imports DTree "~~/src/HOL/Library/Infinite_Set"
@@ -21,7 +21,7 @@
and used: "\<And> n. \<exists> tns. (n,tns) \<in> P"
-subsection{* Tree Basics: frontier, interior, etc. *}
+subsection\<open>Tree Basics: frontier, interior, etc.\<close>
(* Frontier *)
@@ -309,7 +309,7 @@
by (metis inItr.Base subtr_inItr subtr_rootL_in)
-subsection{* The Immediate Subtree Function *}
+subsection\<open>The Immediate Subtree Function\<close>
(* production of: *)
abbreviation "prodOf tr \<equiv> (id \<oplus> root) ` (cont tr)"
@@ -343,7 +343,7 @@
by (metis (lifting) assms image_iff map_sum.simps(2))
-subsection{* Well-Formed Derivation Trees *}
+subsection\<open>Well-Formed Derivation Trees\<close>
hide_const wf
@@ -448,7 +448,7 @@
qed
-subsection{* Default Trees *}
+subsection\<open>Default Trees\<close>
(* Pick a left-hand side of a production for each nonterminal *)
definition S where "S n \<equiv> SOME tns. (n,tns) \<in> P"
@@ -488,7 +488,7 @@
qed
-subsection{* Hereditary Substitution *}
+subsection\<open>Hereditary Substitution\<close>
(* Auxiliary concept: The root-ommiting frontier: *)
definition "inFrr ns tr t \<equiv> \<exists> tr'. Inr tr' \<in> cont tr \<and> inFr ns tr' t"
@@ -679,7 +679,7 @@
end (* context *)
-subsection{* Regular Trees *}
+subsection\<open>Regular Trees\<close>
definition "reg f tr \<equiv> \<forall> tr'. subtr UNIV tr' tr \<longrightarrow> tr' = f (root tr')"
definition "regular tr \<equiv> \<exists> f. reg f tr"
@@ -770,7 +770,7 @@
-subsection {* Paths in a Regular Tree *}
+subsection \<open>Paths in a Regular Tree\<close>
inductive path :: "(N \<Rightarrow> dtree) \<Rightarrow> N list \<Rightarrow> bool" for f where
Base: "path f [n]"
@@ -914,7 +914,7 @@
-subsection{* The Regular Cut of a Tree *}
+subsection\<open>The Regular Cut of a Tree\<close>
context fixes tr0 :: dtree
begin
@@ -1081,7 +1081,7 @@
end (* context *)
-subsection{* Recursive Description of the Regular Tree Frontiers *}
+subsection\<open>Recursive Description of the Regular Tree Frontiers\<close>
lemma regular_inFr:
assumes r: "regular tr" and In: "root tr \<in> ns"
@@ -1130,7 +1130,7 @@
by (simp, metis (lifting) inFr_Ind_minus insert_Diff)
-subsection{* The Generated Languages *}
+subsection\<open>The Generated Languages\<close>
(* The (possibly inifinite tree) generated language *)
definition "L ns n \<equiv> {Fr ns tr | tr. wf tr \<and> root tr = n}"
--- a/src/HOL/Datatype_Examples/Derivation_Trees/Parallel.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Datatype_Examples/Derivation_Trees/Parallel.thy Thu May 26 17:51:22 2016 +0200
@@ -5,7 +5,7 @@
Parallel composition.
*)
-section {* Parallel Composition *}
+section \<open>Parallel Composition\<close>
theory Parallel
imports DTree
@@ -19,7 +19,7 @@
Nplus_comm: "(a::N) + b = b + (a::N)"
and Nplus_assoc: "((a::N) + b) + c = a + (b + c)"
-subsection{* Corecursive Definition of Parallel Composition *}
+subsection\<open>Corecursive Definition of Parallel Composition\<close>
fun par_r where "par_r (tr1,tr2) = root tr1 + root tr2"
fun par_c where
@@ -65,7 +65,7 @@
using Inr_cont_par[of tr1 tr2] unfolding vimage_def by auto
-subsection{* Structural Coinduction Proofs *}
+subsection\<open>Structural Coinduction Proofs\<close>
lemma rel_set_rel_sum_eq[simp]:
"rel_set (rel_sum (op =) \<phi>) A1 A2 \<longleftrightarrow>
--- a/src/HOL/Datatype_Examples/Derivation_Trees/Prelim.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Datatype_Examples/Derivation_Trees/Prelim.thy Thu May 26 17:51:22 2016 +0200
@@ -5,7 +5,7 @@
Preliminaries.
*)
-section {* Preliminaries *}
+section \<open>Preliminaries\<close>
theory Prelim
imports "~~/src/HOL/Library/FSet"
--- a/src/HOL/Datatype_Examples/Koenig.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Datatype_Examples/Koenig.thy Thu May 26 17:51:22 2016 +0200
@@ -6,7 +6,7 @@
Koenig's lemma.
*)
-section {* Koenig's Lemma *}
+section \<open>Koenig's Lemma\<close>
theory Koenig
imports TreeFI "~~/src/HOL/Library/Stream"
--- a/src/HOL/Datatype_Examples/Lambda_Term.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Datatype_Examples/Lambda_Term.thy Thu May 26 17:51:22 2016 +0200
@@ -6,13 +6,13 @@
Lambda-terms.
*)
-section {* Lambda-Terms *}
+section \<open>Lambda-Terms\<close>
theory Lambda_Term
imports "~~/src/HOL/Library/FSet"
begin
-section {* Datatype definition *}
+section \<open>Datatype definition\<close>
datatype 'a trm =
Var 'a |
@@ -21,7 +21,7 @@
Lt "('a \<times> 'a trm) fset" "'a trm"
-subsection {* Example: The set of all variables varsOf and free variables fvarsOf of a term *}
+subsection \<open>Example: The set of all variables varsOf and free variables fvarsOf of a term\<close>
primrec varsOf :: "'a trm \<Rightarrow> 'a set" where
"varsOf (Var a) = {a}"
--- a/src/HOL/Datatype_Examples/Lift_BNF.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Datatype_Examples/Lift_BNF.thy Thu May 26 17:51:22 2016 +0200
@@ -5,7 +5,7 @@
Demonstration of the "lift_bnf" command.
*)
-section {* Demonstration of the \textbf{lift_bnf} Command *}
+section \<open>Demonstration of the \textbf{lift_bnf} Command\<close>
theory Lift_BNF
imports Main
--- a/src/HOL/Datatype_Examples/Misc_Codatatype.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Datatype_Examples/Misc_Codatatype.thy Thu May 26 17:51:22 2016 +0200
@@ -7,7 +7,7 @@
Miscellaneous codatatype definitions.
*)
-section {* Miscellaneous Codatatype Definitions *}
+section \<open>Miscellaneous Codatatype Definitions\<close>
theory Misc_Codatatype
imports "~~/src/HOL/Library/FSet"
--- a/src/HOL/Datatype_Examples/Misc_Datatype.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Datatype_Examples/Misc_Datatype.thy Thu May 26 17:51:22 2016 +0200
@@ -7,7 +7,7 @@
Miscellaneous datatype definitions.
*)
-section {* Miscellaneous Datatype Definitions *}
+section \<open>Miscellaneous Datatype Definitions\<close>
theory Misc_Datatype
imports "~~/src/HOL/Library/Countable" "~~/src/HOL/Library/FSet"
--- a/src/HOL/Datatype_Examples/Misc_Primcorec.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Datatype_Examples/Misc_Primcorec.thy Thu May 26 17:51:22 2016 +0200
@@ -5,7 +5,7 @@
Miscellaneous primitive corecursive function definitions.
*)
-section {* Miscellaneous Primitive Corecursive Function Definitions *}
+section \<open>Miscellaneous Primitive Corecursive Function Definitions\<close>
theory Misc_Primcorec
imports Misc_Codatatype
--- a/src/HOL/Datatype_Examples/Misc_Primrec.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Datatype_Examples/Misc_Primrec.thy Thu May 26 17:51:22 2016 +0200
@@ -5,7 +5,7 @@
Miscellaneous primitive recursive function definitions.
*)
-section {* Miscellaneous Primitive Recursive Function Definitions *}
+section \<open>Miscellaneous Primitive Recursive Function Definitions\<close>
theory Misc_Primrec
imports Misc_Datatype
--- a/src/HOL/Datatype_Examples/Process.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Datatype_Examples/Process.thy Thu May 26 17:51:22 2016 +0200
@@ -5,7 +5,7 @@
Processes.
*)
-section {* Processes *}
+section \<open>Processes\<close>
theory Process
imports "~~/src/HOL/Library/Stream"
@@ -17,7 +17,7 @@
(* Read: prefix of, continuation of, choice 1 of, choice 2 of *)
-subsection {* Basic properties *}
+subsection \<open>Basic properties\<close>
(* Constructors versus discriminators *)
theorem isAction_isChoice:
@@ -28,7 +28,7 @@
by (cases rule: process.exhaust[of p]) auto
-subsection{* Coinduction *}
+subsection\<open>Coinduction\<close>
theorem process_coind[elim, consumes 1, case_names iss Action Choice, induct pred: "HOL.eq"]:
assumes phi: "\<phi> p p'" and
@@ -50,19 +50,19 @@
by (coinduct rule: process.coinduct_strong) (metis process.collapse(1,2) process.disc(3))
-subsection {* Coiteration (unfold) *}
+subsection \<open>Coiteration (unfold)\<close>
-section{* Coinductive definition of the notion of trace *}
+section\<open>Coinductive definition of the notion of trace\<close>
coinductive trace where
"trace p as \<Longrightarrow> trace (Action a p) (a ## as)"
|
"trace p as \<or> trace q as \<Longrightarrow> trace (Choice p q) as"
-section{* Examples of corecursive definitions: *}
+section\<open>Examples of corecursive definitions:\<close>
-subsection{* Single-guard fixpoint definition *}
+subsection\<open>Single-guard fixpoint definition\<close>
primcorec BX where
"isAction BX"
@@ -70,7 +70,7 @@
| "contOf BX = BX"
-subsection{* Multi-guard fixpoint definitions, simulated with auxiliary arguments *}
+subsection\<open>Multi-guard fixpoint definitions, simulated with auxiliary arguments\<close>
datatype x_y_ax = x | y | ax
@@ -94,7 +94,7 @@
-section{* Case study: Multi-guard fixpoint definitions, without auxiliary arguments *}
+section\<open>Case study: Multi-guard fixpoint definitions, without auxiliary arguments\<close>
hide_const x y ax X Y AX
--- a/src/HOL/Datatype_Examples/Stream_Processor.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Datatype_Examples/Stream_Processor.thy Thu May 26 17:51:22 2016 +0200
@@ -6,7 +6,7 @@
Stream processors---a syntactic representation of continuous functions on streams.
*)
-section {* Stream Processors---A Syntactic Representation of Continuous Functions on Streams *}
+section \<open>Stream Processors---A Syntactic Representation of Continuous Functions on Streams\<close>
theory Stream_Processor
imports "~~/src/HOL/Library/Stream" "~~/src/HOL/Library/BNF_Axiomatization"
@@ -15,7 +15,7 @@
declare [[typedef_overloaded]]
-section {* Continuous Functions on Streams *}
+section \<open>Continuous Functions on Streams\<close>
datatype ('a, 'b, 'c) sp\<^sub>\<mu> =
Get "'a \<Rightarrow> ('a, 'b, 'c) sp\<^sub>\<mu>"
@@ -37,10 +37,10 @@
lemma run\<^sub>\<nu>_copy: "run\<^sub>\<nu> copy s = s"
by (coinduction arbitrary: s) simp
-text {*
+text \<open>
To use the function package for the definition of composition the
wellfoundedness of the subtree relation needs to be proved first.
-*}
+\<close>
definition "sub \<equiv> {(f a, Get f) | a f. True}"
@@ -88,7 +88,7 @@
qed
qed
-text {* Alternative definition of composition using primrec instead of function *}
+text \<open>Alternative definition of composition using primrec instead of function\<close>
primrec sp\<^sub>\<mu>_comp2R where
"sp\<^sub>\<mu>_comp2R f (Put b sp) = f b (out sp)"
@@ -125,7 +125,7 @@
qed
qed
-text {* The two definitions are equivalent *}
+text \<open>The two definitions are equivalent\<close>
lemma sp\<^sub>\<mu>_comp_sp\<^sub>\<mu>_comp2[simp]: "sp o\<^sub>\<mu> sp' = sp o\<^sup>*\<^sub>\<mu> sp'"
by (induct sp sp' rule: sp\<^sub>\<mu>_comp.induct) auto
@@ -134,7 +134,7 @@
lemma sp\<^sub>\<mu>_rel_map_map[unfolded vimage2p_def, simp]:
"rel_sp\<^sub>\<mu> R1 R2 (map_sp\<^sub>\<mu> f1 f2 sp) (map_sp\<^sub>\<mu> g1 g2 sp') =
rel_sp\<^sub>\<mu> (BNF_Def.vimage2p f1 g1 R1) (BNF_Def.vimage2p f2 g2 R2) sp sp'"
-by (tactic {*
+by (tactic \<open>
let
val ks = 1 upto 2;
val ctxt = @{context};
@@ -157,7 +157,7 @@
hyp_subst_tac ctxt,
assume_tac ctxt])
end
-*})
+\<close>)
lemma sp\<^sub>\<mu>_rel_self: "\<lbrakk>op = \<le> R1; op = \<le> R2\<rbrakk> \<Longrightarrow> rel_sp\<^sub>\<mu> R1 R2 x x"
by (erule (1) predicate2D[OF sp\<^sub>\<mu>.rel_mono]) (simp only: sp\<^sub>\<mu>.rel_eq)
@@ -166,7 +166,7 @@
by (coinduction arbitrary: sp sp') (auto intro!: sp\<^sub>\<mu>_rel_self)
-section {* Generalization to an Arbitrary BNF as Codomain *}
+section \<open>Generalization to an Arbitrary BNF as Codomain\<close>
bnf_axiomatization ('a, 'b) F for map: F
--- a/src/HOL/Datatype_Examples/TreeFI.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Datatype_Examples/TreeFI.thy Thu May 26 17:51:22 2016 +0200
@@ -6,7 +6,7 @@
Finitely branching possibly infinite trees.
*)
-section {* Finitely Branching Possibly Infinite Trees *}
+section \<open>Finitely Branching Possibly Infinite Trees\<close>
theory TreeFI
imports Main
--- a/src/HOL/Datatype_Examples/TreeFsetI.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Datatype_Examples/TreeFsetI.thy Thu May 26 17:51:22 2016 +0200
@@ -6,7 +6,7 @@
Finitely branching possibly infinite trees, with sets of children.
*)
-section {* Finitely Branching Possibly Infinite Trees, with Sets of Children *}
+section \<open>Finitely Branching Possibly Infinite Trees, with Sets of Children\<close>
theory TreeFsetI
imports "~~/src/HOL/Library/FSet"
--- a/src/HOL/IMPP/Com.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/IMPP/Com.thy Thu May 26 17:51:22 2016 +0200
@@ -2,7 +2,7 @@
Author: David von Oheimb (based on a theory by Tobias Nipkow et al), TUM
*)
-section {* Semantics of arithmetic and boolean expressions, Syntax of commands *}
+section \<open>Semantics of arithmetic and boolean expressions, Syntax of commands\<close>
theory Com
imports Main
@@ -83,10 +83,10 @@
"WT_bodies = (!(pn,b):set bodies. WT b)"
-ML {*
+ML \<open>
fun make_imp_tac ctxt =
EVERY' [resolve_tac ctxt [mp], fn i => assume_tac ctxt (i + 1), eresolve_tac ctxt [thin_rl]]
-*}
+\<close>
lemma finite_dom_body: "finite (dom body)"
apply (unfold body_def)
--- a/src/HOL/IMPP/EvenOdd.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/IMPP/EvenOdd.thy Thu May 26 17:51:22 2016 +0200
@@ -2,7 +2,7 @@
Author: David von Oheimb, TUM
*)
-section {* Example of mutually recursive procedures verified with Hoare logic *}
+section \<open>Example of mutually recursive procedures verified with Hoare logic\<close>
theory EvenOdd
imports Main Misc
--- a/src/HOL/IMPP/Hoare.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/IMPP/Hoare.thy Thu May 26 17:51:22 2016 +0200
@@ -3,18 +3,18 @@
Copyright 1999 TUM
*)
-section {* Inductive definition of Hoare logic for partial correctness *}
+section \<open>Inductive definition of Hoare logic for partial correctness\<close>
theory Hoare
imports Natural
begin
-text {*
+text \<open>
Completeness is taken relative to completeness of the underlying logic.
Two versions of completeness proof: nested single recursion
vs. simultaneous recursion in call rule
-*}
+\<close>
type_synonym 'a assn = "'a => state => bool"
translations
@@ -105,7 +105,7 @@
X:=CALL pn(a) .{Q}"
-section {* Soundness and relative completeness of Hoare rules wrt operational semantics *}
+section \<open>Soundness and relative completeness of Hoare rules wrt operational semantics\<close>
lemma single_stateE:
"state_not_singleton ==> !t. (!s::state. s = t) --> False"
@@ -209,7 +209,7 @@
*)
lemma thin [rule_format]: "G'||-ts ==> !G. G' <= G --> G||-ts"
apply (erule hoare_derivs.induct)
-apply (tactic {* ALLGOALS (EVERY'[clarify_tac @{context}, REPEAT o smp_tac @{context} 1]) *})
+apply (tactic \<open>ALLGOALS (EVERY'[clarify_tac @{context}, REPEAT o smp_tac @{context} 1])\<close>)
apply (rule hoare_derivs.empty)
apply (erule (1) hoare_derivs.insert)
apply (fast intro: hoare_derivs.asm)
@@ -218,7 +218,7 @@
apply (rule hoare_derivs.conseq, intro strip, tactic "smp_tac @{context} 2 1", clarify, tactic "smp_tac @{context} 1 1",rule exI, rule exI, erule (1) conjI)
prefer 7
apply (rule_tac hoare_derivs.Body, drule_tac spec, erule_tac mp, fast)
-apply (tactic {* ALLGOALS (resolve_tac @{context} ((funpow 5 tl) @{thms hoare_derivs.intros}) THEN_ALL_NEW (fast_tac @{context})) *})
+apply (tactic \<open>ALLGOALS (resolve_tac @{context} ((funpow 5 tl) @{thms hoare_derivs.intros}) THEN_ALL_NEW (fast_tac @{context}))\<close>)
done
lemma weak_Body: "G|-{P}. the (body pn) .{Q} ==> G|-{P}. BODY pn .{Q}"
@@ -278,19 +278,19 @@
lemma hoare_sound: "G||-ts ==> G||=ts"
apply (erule hoare_derivs.induct)
-apply (tactic {* TRYALL (eresolve_tac @{context} [@{thm Loop_sound_lemma}, @{thm Body_sound_lemma}] THEN_ALL_NEW assume_tac @{context}) *})
+apply (tactic \<open>TRYALL (eresolve_tac @{context} [@{thm Loop_sound_lemma}, @{thm Body_sound_lemma}] THEN_ALL_NEW assume_tac @{context})\<close>)
apply (unfold hoare_valids_def)
apply blast
apply blast
apply (blast) (* asm *)
apply (blast) (* cut *)
apply (blast) (* weaken *)
-apply (tactic {* ALLGOALS (EVERY'
+apply (tactic \<open>ALLGOALS (EVERY'
[REPEAT o Rule_Insts.thin_tac @{context} "hoare_derivs _ _" [],
- simp_tac @{context}, clarify_tac @{context}, REPEAT o smp_tac @{context} 1]) *})
+ simp_tac @{context}, clarify_tac @{context}, REPEAT o smp_tac @{context} 1])\<close>)
apply (simp_all (no_asm_use) add: triple_valid_def2)
apply (intro strip, tactic "smp_tac @{context} 2 1", blast) (* conseq *)
-apply (tactic {* ALLGOALS (clarsimp_tac @{context}) *}) (* Skip, Ass, Local *)
+apply (tactic \<open>ALLGOALS (clarsimp_tac @{context})\<close>) (* Skip, Ass, Local *)
prefer 3 apply (force) (* Call *)
apply (erule_tac [2] evaln_elim_cases) (* If *)
apply blast+
@@ -335,24 +335,24 @@
lemma MGF_lemma1 [rule_format (no_asm)]: "state_not_singleton ==>
!pn:dom body. G|-{=}.BODY pn.{->} ==> WT c --> G|-{=}.c.{->}"
apply (induct_tac c)
-apply (tactic {* ALLGOALS (clarsimp_tac @{context}) *})
+apply (tactic \<open>ALLGOALS (clarsimp_tac @{context})\<close>)
prefer 7 apply (fast intro: domI)
apply (erule_tac [6] MGT_alternD)
apply (unfold MGT_def)
apply (drule_tac [7] bspec, erule_tac [7] domI)
-apply (rule_tac [7] escape, tactic {* clarsimp_tac @{context} 7 *},
+apply (rule_tac [7] escape, tactic \<open>clarsimp_tac @{context} 7\<close>,
rename_tac [7] "fun" y Z,
rule_tac [7] P1 = "%Z' s. s= (setlocs Z newlocs) [Loc Arg ::= fun Z]" in hoare_derivs.Call [THEN conseq1], erule_tac [7] conseq12)
apply (erule_tac [!] thin_rl)
apply (rule hoare_derivs.Skip [THEN conseq2])
apply (rule_tac [2] hoare_derivs.Ass [THEN conseq1])
-apply (rule_tac [3] escape, tactic {* clarsimp_tac @{context} 3 *},
+apply (rule_tac [3] escape, tactic \<open>clarsimp_tac @{context} 3\<close>,
rename_tac [3] loc "fun" y Z,
rule_tac [3] P1 = "%Z' s. s= (Z[Loc loc::=fun Z])" in hoare_derivs.Local [THEN conseq1],
erule_tac [3] conseq12)
apply (erule_tac [5] hoare_derivs.Comp, erule_tac [5] conseq12)
-apply (tactic {* (resolve_tac @{context} @{thms hoare_derivs.If} THEN_ALL_NEW
- eresolve_tac @{context} @{thms conseq12}) 6 *})
+apply (tactic \<open>(resolve_tac @{context} @{thms hoare_derivs.If} THEN_ALL_NEW
+ eresolve_tac @{context} @{thms conseq12}) 6\<close>)
apply (rule_tac [8] hoare_derivs.Loop [THEN conseq2], erule_tac [8] conseq12)
apply auto
done
@@ -367,7 +367,7 @@
shows "finite U ==> uG = mgt_call`U ==>
!G. G <= uG --> n <= card uG --> card G = card uG - n --> (!c. wt c --> P G {mgt c})"
apply (induct_tac n)
-apply (tactic {* ALLGOALS (clarsimp_tac @{context}) *})
+apply (tactic \<open>ALLGOALS (clarsimp_tac @{context})\<close>)
apply (subgoal_tac "G = mgt_call ` U")
prefer 2
apply (simp add: card_seteq)
--- a/src/HOL/IMPP/Misc.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/IMPP/Misc.thy Thu May 26 17:51:22 2016 +0200
@@ -2,7 +2,7 @@
Author: David von Oheimb, TUM
*)
-section {* Several examples for Hoare logic *}
+section \<open>Several examples for Hoare logic\<close>
theory Misc
imports Hoare
--- a/src/HOL/IMPP/Natural.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/IMPP/Natural.thy Thu May 26 17:51:22 2016 +0200
@@ -2,7 +2,7 @@
Author: David von Oheimb (based on a theory by Tobias Nipkow et al), TUM
*)
-section {* Natural semantics of commands *}
+section \<open>Natural semantics of commands\<close>
theory Natural
imports Com
@@ -111,9 +111,9 @@
lemma evaln_evalc: "<c,s> -n-> t ==> <c,s> -c-> t"
apply (erule evaln.induct)
-apply (tactic {*
+apply (tactic \<open>
ALLGOALS (resolve_tac @{context} @{thms evalc.intros} THEN_ALL_NEW assume_tac @{context})
-*})
+\<close>)
done
lemma Suc_le_D_lemma: "[| Suc n <= m'; (!!m. n <= m ==> P (Suc m)) |] ==> P m'"
@@ -139,12 +139,12 @@
lemma evalc_evaln: "<c,s> -c-> t ==> ? n. <c,s> -n-> t"
apply (erule evalc.induct)
-apply (tactic {* ALLGOALS (REPEAT o eresolve_tac @{context} [exE]) *})
-apply (tactic {* TRYALL (EVERY' [dresolve_tac @{context} @{thms evaln_max2}, assume_tac @{context},
- REPEAT o eresolve_tac @{context} [exE, conjE]]) *})
+apply (tactic \<open>ALLGOALS (REPEAT o eresolve_tac @{context} [exE])\<close>)
+apply (tactic \<open>TRYALL (EVERY' [dresolve_tac @{context} @{thms evaln_max2}, assume_tac @{context},
+ REPEAT o eresolve_tac @{context} [exE, conjE]])\<close>)
apply (tactic
- {* ALLGOALS (resolve_tac @{context} [exI] THEN'
- resolve_tac @{context} @{thms evaln.intros} THEN_ALL_NEW assume_tac @{context}) *})
+ \<open>ALLGOALS (resolve_tac @{context} [exI] THEN'
+ resolve_tac @{context} @{thms evaln.intros} THEN_ALL_NEW assume_tac @{context})\<close>)
done
lemma eval_eq: "<c,s> -c-> t = (? n. <c,s> -n-> t)"
--- a/src/HOL/IOA/Asig.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/IOA/Asig.thy Thu May 26 17:51:22 2016 +0200
@@ -3,7 +3,7 @@
Copyright 1994 TU Muenchen
*)
-section {* Action signatures *}
+section \<open>Action signatures\<close>
theory Asig
imports Main
--- a/src/HOL/IOA/IOA.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/IOA/IOA.thy Thu May 26 17:51:22 2016 +0200
@@ -3,7 +3,7 @@
Copyright 1994 TU Muenchen
*)
-section {* The I/O automata of Lynch and Tuttle *}
+section \<open>The I/O automata of Lynch and Tuttle\<close>
theory IOA
imports Asig
--- a/src/HOL/IOA/Solve.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/IOA/Solve.thy Thu May 26 17:51:22 2016 +0200
@@ -3,7 +3,7 @@
Copyright 1994 TU Muenchen
*)
-section {* Weak possibilities mapping (abstraction) *}
+section \<open>Weak possibilities mapping (abstraction)\<close>
theory Solve
imports IOA
@@ -144,9 +144,9 @@
prefer 2
apply force
apply (simp (no_asm) add: conj_disj_distribR cong add: conj_cong split add: if_split)
- apply (tactic {*
+ apply (tactic \<open>
REPEAT((resolve_tac @{context} [conjI, impI] 1 ORELSE eresolve_tac @{context} [conjE] 1) THEN
- asm_full_simp_tac(@{context} addsimps [@{thm comp1_reachable}, @{thm comp2_reachable}]) 1) *})
+ asm_full_simp_tac(@{context} addsimps [@{thm comp1_reachable}, @{thm comp2_reachable}]) 1)\<close>)
done
--- a/src/HOL/Imperative_HOL/Array.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Imperative_HOL/Array.thy Thu May 26 17:51:22 2016 +0200
@@ -2,13 +2,13 @@
Author: John Matthews, Galois Connections; Alexander Krauss, Lukas Bulwahn & Florian Haftmann, TU Muenchen
*)
-section {* Monadic arrays *}
+section \<open>Monadic arrays\<close>
theory Array
imports Heap_Monad
begin
-subsection {* Primitives *}
+subsection \<open>Primitives\<close>
definition present :: "heap \<Rightarrow> 'a::heap array \<Rightarrow> bool" where
"present h a \<longleftrightarrow> addr_of_array a < lim h"
@@ -36,7 +36,7 @@
"r =!!= s \<longleftrightarrow> TYPEREP('a) \<noteq> TYPEREP('b) \<or> addr_of_array r \<noteq> addr_of_array s"
-subsection {* Monad operations *}
+subsection \<open>Monad operations\<close>
definition new :: "nat \<Rightarrow> 'a::heap \<Rightarrow> 'a array Heap" where
[code del]: "new n x = Heap_Monad.heap (alloc (replicate n x))"
@@ -70,12 +70,12 @@
[code del]: "freeze a = Heap_Monad.tap (\<lambda>h. get h a)"
-subsection {* Properties *}
+subsection \<open>Properties\<close>
-text {* FIXME: Does there exist a "canonical" array axiomatisation in
-the literature? *}
+text \<open>FIXME: Does there exist a "canonical" array axiomatisation in
+the literature?\<close>
-text {* Primitives *}
+text \<open>Primitives\<close>
lemma noteq_sym: "a =!!= b \<Longrightarrow> b =!!= a"
and unequal [simp]: "a \<noteq> a' \<longleftrightarrow> a =!!= a'"
@@ -160,7 +160,7 @@
by (simp add: present_def alloc_def Let_def)
-text {* Monad operations *}
+text \<open>Monad operations\<close>
lemma execute_new [execute_simps]:
"execute (new n x) h = Some (alloc (replicate n x) h)"
@@ -352,9 +352,9 @@
hide_const (open) present get set alloc length update noteq new of_list make len nth upd map_entry swap freeze
-subsection {* Code generator setup *}
+subsection \<open>Code generator setup\<close>
-subsubsection {* Logical intermediate layer *}
+subsubsection \<open>Logical intermediate layer\<close>
definition new' where
[code del]: "new' = Array.new o nat_of_integer"
@@ -439,7 +439,7 @@
hide_const (open) new' make' len' nth' upd'
-text {* SML *}
+text \<open>SML\<close>
code_printing type_constructor array \<rightharpoonup> (SML) "_/ array"
code_printing constant Array \<rightharpoonup> (SML) "raise/ (Fail/ \"bare Array\")"
@@ -454,7 +454,7 @@
code_reserved SML Array
-text {* OCaml *}
+text \<open>OCaml\<close>
code_printing type_constructor array \<rightharpoonup> (OCaml) "_/ array"
code_printing constant Array \<rightharpoonup> (OCaml) "failwith/ \"bare Array\""
@@ -470,7 +470,7 @@
code_reserved OCaml Array
-text {* Haskell *}
+text \<open>Haskell\<close>
code_printing type_constructor array \<rightharpoonup> (Haskell) "Heap.STArray/ Heap.RealWorld/ _"
code_printing constant Array \<rightharpoonup> (Haskell) "error/ \"bare Array\""
@@ -484,7 +484,7 @@
code_printing class_instance array :: HOL.equal \<rightharpoonup> (Haskell) -
-text {* Scala *}
+text \<open>Scala\<close>
code_printing type_constructor array \<rightharpoonup> (Scala) "!collection.mutable.ArraySeq[_]"
code_printing constant Array \<rightharpoonup> (Scala) "!sys.error(\"bare Array\")"
--- a/src/HOL/Imperative_HOL/Heap.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Imperative_HOL/Heap.thy Thu May 26 17:51:22 2016 +0200
@@ -2,15 +2,15 @@
Author: John Matthews, Galois Connections; Alexander Krauss, TU Muenchen
*)
-section {* A polymorphic heap based on cantor encodings *}
+section \<open>A polymorphic heap based on cantor encodings\<close>
theory Heap
imports Main "~~/src/HOL/Library/Countable"
begin
-subsection {* Representable types *}
+subsection \<open>Representable types\<close>
-text {* The type class of representable types *}
+text \<open>The type class of representable types\<close>
class heap = typerep + countable
@@ -35,15 +35,15 @@
instance typerep :: heap ..
-subsection {* A polymorphic heap with dynamic arrays and references *}
+subsection \<open>A polymorphic heap with dynamic arrays and references\<close>
-text {*
+text \<open>
References and arrays are developed in parallel,
but keeping them separate makes some later proofs simpler.
-*}
+\<close>
-type_synonym addr = nat -- "untyped heap references"
-type_synonym heap_rep = nat -- "representable values"
+type_synonym addr = nat \<comment> "untyped heap references"
+type_synonym heap_rep = nat \<comment> "representable values"
record heap =
arrays :: "typerep \<Rightarrow> addr \<Rightarrow> heap_rep list"
@@ -53,8 +53,8 @@
definition empty :: heap where
"empty = \<lparr>arrays = (\<lambda>_ _. []), refs = (\<lambda>_ _. 0), lim = 0\<rparr>"
-datatype 'a array = Array addr -- "note the phantom type 'a"
-datatype 'a ref = Ref addr -- "note the phantom type 'a"
+datatype 'a array = Array addr \<comment> "note the phantom type 'a"
+datatype 'a ref = Ref addr \<comment> "note the phantom type 'a"
primrec addr_of_array :: "'a array \<Rightarrow> addr" where
"addr_of_array (Array x) = x"
@@ -76,14 +76,14 @@
instance ref :: (type) countable
by (rule countable_classI [of addr_of_ref]) simp
-text {* Syntactic convenience *}
+text \<open>Syntactic convenience\<close>
-setup {*
+setup \<open>
Sign.add_const_constraint (@{const_name Array}, SOME @{typ "nat \<Rightarrow> 'a::heap array"})
#> Sign.add_const_constraint (@{const_name Ref}, SOME @{typ "nat \<Rightarrow> 'a::heap ref"})
#> Sign.add_const_constraint (@{const_name addr_of_array}, SOME @{typ "'a::heap array \<Rightarrow> nat"})
#> Sign.add_const_constraint (@{const_name addr_of_ref}, SOME @{typ "'a::heap ref \<Rightarrow> nat"})
-*}
+\<close>
hide_const (open) empty
--- a/src/HOL/Imperative_HOL/Heap_Monad.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Imperative_HOL/Heap_Monad.thy Thu May 26 17:51:22 2016 +0200
@@ -2,7 +2,7 @@
Author: John Matthews, Galois Connections; Alexander Krauss, Lukas Bulwahn & Florian Haftmann, TU Muenchen
*)
-section {* A monad with a polymorphic heap and primitive reasoning infrastructure *}
+section \<open>A monad with a polymorphic heap and primitive reasoning infrastructure\<close>
theory Heap_Monad
imports
@@ -10,12 +10,12 @@
"~~/src/HOL/Library/Monad_Syntax"
begin
-subsection {* The monad *}
+subsection \<open>The monad\<close>
-subsubsection {* Monad construction *}
+subsubsection \<open>Monad construction\<close>
-text {* Monadic heap actions either produce values
- and transform the heap, or fail *}
+text \<open>Monadic heap actions either produce values
+ and transform the heap, or fail\<close>
datatype 'a Heap = Heap "heap \<Rightarrow> ('a \<times> heap) option"
lemma [code, code del]:
@@ -46,7 +46,7 @@
by (simp add: Let_def)
-subsubsection {* Specialised lifters *}
+subsubsection \<open>Specialised lifters\<close>
definition tap :: "(heap \<Rightarrow> 'a) \<Rightarrow> 'a Heap" where
[code del]: "tap f = Heap (\<lambda>h. Some (f h, h))"
@@ -71,7 +71,7 @@
by (simp_all add: guard_def)
-subsubsection {* Predicate classifying successful computations *}
+subsubsection \<open>Predicate classifying successful computations\<close>
definition success :: "'a Heap \<Rightarrow> heap \<Rightarrow> bool" where
"success f h \<longleftrightarrow> execute f h \<noteq> None"
@@ -109,13 +109,13 @@
by (simp add: success_def)
-subsubsection {* Predicate for a simple relational calculus *}
+subsubsection \<open>Predicate for a simple relational calculus\<close>
-text {*
- The @{text effect} predicate states that when a computation @{text c}
- runs with the heap @{text h} will result in return value @{text r}
- and a heap @{text "h'"}, i.e.~no exception occurs.
-*}
+text \<open>
+ The \<open>effect\<close> predicate states that when a computation \<open>c\<close>
+ runs with the heap \<open>h\<close> will result in return value \<open>r\<close>
+ and a heap \<open>h'\<close>, i.e.~no exception occurs.
+\<close>
definition effect :: "'a Heap \<Rightarrow> heap \<Rightarrow> heap \<Rightarrow> 'a \<Rightarrow> bool" where
effect_def: "effect c h h' r \<longleftrightarrow> execute c h = Some (r, h')"
@@ -210,7 +210,7 @@
(auto simp add: execute_simps elim!: successE, cases "P h", auto simp add: execute_simps)
-subsubsection {* Monad combinators *}
+subsubsection \<open>Monad combinators\<close>
definition return :: "'a \<Rightarrow> 'a Heap" where
[code del]: "return x = heap (Pair x)"
@@ -232,7 +232,7 @@
obtains "r = x" "h' = h"
using assms by (rule effectE) (simp add: execute_simps)
-definition raise :: "string \<Rightarrow> 'a Heap" where -- {* the string is just decoration *}
+definition raise :: "string \<Rightarrow> 'a Heap" where \<comment> \<open>the string is just decoration\<close>
[code del]: "raise s = Heap (\<lambda>_. None)"
lemma execute_raise [execute_simps]:
@@ -306,9 +306,9 @@
by (rule Heap_eqI) (simp add: execute_simps)
-subsection {* Generic combinators *}
+subsection \<open>Generic combinators\<close>
-subsubsection {* Assertions *}
+subsubsection \<open>Assertions\<close>
definition assert :: "('a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'a Heap" where
"assert P x = (if P x then return x else raise ''assert'')"
@@ -338,7 +338,7 @@
by (rule Heap_eqI) (insert assms, simp add: assert_def)
-subsubsection {* Plain lifting *}
+subsubsection \<open>Plain lifting\<close>
definition lift :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b Heap" where
"lift f = return o f"
@@ -352,7 +352,7 @@
by (simp add: lift_def comp_def)
-subsubsection {* Iteration -- warning: this is rarely useful! *}
+subsubsection \<open>Iteration -- warning: this is rarely useful!\<close>
primrec fold_map :: "('a \<Rightarrow> 'b Heap) \<Rightarrow> 'a list \<Rightarrow> 'b list Heap" where
"fold_map f [] = return []"
@@ -382,7 +382,7 @@
qed
-subsection {* Partial function definition setup *}
+subsection \<open>Partial function definition setup\<close>
definition Heap_ord :: "'a Heap \<Rightarrow> 'a Heap \<Rightarrow> bool" where
"Heap_ord = img_ord execute (fun_ord option_ord)"
@@ -455,9 +455,9 @@
unfolding effect_def execute.simps
by blast
-declaration {* Partial_Function.init "heap" @{term heap.fixp_fun}
+declaration \<open>Partial_Function.init "heap" @{term heap.fixp_fun}
@{term heap.mono_body} @{thm heap.fixp_rule_uc} @{thm heap.fixp_induct_uc}
- (SOME @{thm fixp_induct_heap}) *}
+ (SOME @{thm fixp_induct_heap})\<close>
abbreviation "mono_Heap \<equiv> monotone (fun_ord Heap_ord) Heap_ord"
@@ -516,9 +516,9 @@
qed
-subsection {* Code generator setup *}
+subsection \<open>Code generator setup\<close>
-subsubsection {* Logical intermediate layer *}
+subsubsection \<open>Logical intermediate layer\<close>
definition raise' :: "String.literal \<Rightarrow> 'a Heap" where
[code del]: "raise' s = raise (String.explode s)"
@@ -530,10 +530,10 @@
"raise s = raise' (STR s)"
unfolding raise'_def by (simp add: STR_inverse)
-code_datatype raise' -- {* avoid @{const "Heap"} formally *}
+code_datatype raise' \<comment> \<open>avoid @{const "Heap"} formally\<close>
-subsubsection {* SML and OCaml *}
+subsubsection \<open>SML and OCaml\<close>
code_printing type_constructor Heap \<rightharpoonup> (SML) "(unit/ ->/ _)"
code_printing constant bind \<rightharpoonup> (SML) "!(fn/ f'_/ =>/ fn/ ()/ =>/ f'_/ (_/ ())/ ())"
@@ -546,12 +546,12 @@
code_printing constant Heap_Monad.raise' \<rightharpoonup> (OCaml) "failwith"
-subsubsection {* Haskell *}
+subsubsection \<open>Haskell\<close>
-text {* Adaption layer *}
+text \<open>Adaption layer\<close>
code_printing code_module "Heap" \<rightharpoonup> (Haskell)
-{*import qualified Control.Monad;
+\<open>import qualified Control.Monad;
import qualified Control.Monad.ST;
import qualified Data.STRef;
import qualified Data.Array.ST;
@@ -581,11 +581,11 @@
readArray = Data.Array.ST.readArray;
writeArray :: STArray s a -> Integer -> a -> ST s ();
-writeArray = Data.Array.ST.writeArray;*}
+writeArray = Data.Array.ST.writeArray;\<close>
code_reserved Haskell Heap
-text {* Monad *}
+text \<open>Monad\<close>
code_printing type_constructor Heap \<rightharpoonup> (Haskell) "Heap.ST/ Heap.RealWorld/ _"
code_monad bind Haskell
@@ -593,10 +593,10 @@
code_printing constant Heap_Monad.raise' \<rightharpoonup> (Haskell) "error"
-subsubsection {* Scala *}
+subsubsection \<open>Scala\<close>
code_printing code_module "Heap" \<rightharpoonup> (Scala)
-{*object Heap {
+\<open>object Heap {
def bind[A, B](f: Unit => A, g: A => Unit => B): Unit => B = (_: Unit) => g (f ()) ()
}
@@ -628,7 +628,7 @@
def freeze[A](a: ArraySeq[A]): List[A] =
a.toList
}
-*}
+\<close>
code_reserved Scala Heap Ref Array
@@ -638,9 +638,9 @@
code_printing constant Heap_Monad.raise' \<rightharpoonup> (Scala) "!sys.error((_))"
-subsubsection {* Target variants with less units *}
+subsubsection \<open>Target variants with less units\<close>
-setup {*
+setup \<open>
let
@@ -703,7 +703,7 @@
end
-*}
+\<close>
hide_const (open) Heap heap guard raise' fold_map
--- a/src/HOL/Imperative_HOL/Imperative_HOL.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Imperative_HOL/Imperative_HOL.thy Thu May 26 17:51:22 2016 +0200
@@ -2,7 +2,7 @@
Author: John Matthews, Galois Connections; Alexander Krauss, Lukas Bulwahn & Florian Haftmann, TU Muenchen
*)
-section {* Entry point into monadic imperative HOL *}
+section \<open>Entry point into monadic imperative HOL\<close>
theory Imperative_HOL
imports Array Ref
--- a/src/HOL/Imperative_HOL/Imperative_HOL_ex.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Imperative_HOL/Imperative_HOL_ex.thy Thu May 26 17:51:22 2016 +0200
@@ -3,7 +3,7 @@
Alexander Krauss, Lukas Bulwahn & Florian Haftmann, TU Muenchen
*)
-section {* Monadic imperative HOL with examples *}
+section \<open>Monadic imperative HOL with examples\<close>
theory Imperative_HOL_ex
imports Imperative_HOL Overview
--- a/src/HOL/Imperative_HOL/Overview.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Imperative_HOL/Overview.thy Thu May 26 17:51:22 2016 +0200
@@ -17,9 +17,9 @@
"_constrain" :: "prop' => type => prop'" ("_ :: _" [4, 0] 3)
(*>*)
-text {*
- @{text "Imperative HOL"} is a leightweight framework for reasoning
- about imperative data structures in @{text "Isabelle/HOL"}
+text \<open>
+ \<open>Imperative HOL\<close> is a leightweight framework for reasoning
+ about imperative data structures in \<open>Isabelle/HOL\<close>
@{cite "Nipkow-et-al:2002:tutorial"}. Its basic ideas are described in
@{cite "Bulwahn-et-al:2008:imp_HOL"}. However their concrete
realisation has changed since, due to both extensions and
@@ -27,16 +27,16 @@
\qt{as it is} by now. It focusses on the user-view, less on matters
of construction. For details study of the theory sources is
encouraged.
-*}
+\<close>
-section {* A polymorphic heap inside a monad *}
+section \<open>A polymorphic heap inside a monad\<close>
-text {*
+text \<open>
Heaps (@{type heap}) can be populated by values of class @{class
heap}; HOL's default types are already instantiated to class @{class
heap}. Class @{class heap} is a subclass of @{class countable}; see
- theory @{text Countable} for ways to instantiate types as @{class countable}.
+ theory \<open>Countable\<close> for ways to instantiate types as @{class countable}.
The heap is wrapped up in a monad @{typ "'a Heap"} by means of the
following specification:
@@ -53,11 +53,11 @@
\end{quote}
This allows for equational reasoning about monadic expressions; the
- fact collection @{text execute_simps} contains appropriate rewrites
+ fact collection \<open>execute_simps\<close> contains appropriate rewrites
for all fundamental operations.
Primitive fine-granular control over heaps is available through rule
- @{text Heap_cases}:
+ \<open>Heap_cases\<close>:
\begin{quote}
@{thm [break] Heap_cases [no_vars]}
@@ -81,22 +81,21 @@
@{term_type assert} \\
@{thm assert_def [no_vars]}
\end{quote}
-*}
+\<close>
-section {* Relational reasoning about @{type Heap} expressions *}
+section \<open>Relational reasoning about @{type Heap} expressions\<close>
-text {*
+text \<open>
To establish correctness of imperative programs, predicate
\begin{quote}
@{term_type effect}
\end{quote}
- provides a simple relational calculus. Primitive rules are @{text
- effectI} and @{text effectE}, rules appropriate for reasoning about
- imperative operations are available in the @{text effect_intros} and
- @{text effect_elims} fact collections.
+ provides a simple relational calculus. Primitive rules are \<open>effectI\<close> and \<open>effectE\<close>, rules appropriate for reasoning about
+ imperative operations are available in the \<open>effect_intros\<close> and
+ \<open>effect_elims\<close> fact collections.
Often non-failure of imperative computations does not depend
on the heap at all; reasoning then can be easier using predicate
@@ -106,18 +105,16 @@
\end{quote}
Introduction rules for @{const success} are available in the
- @{text success_intro} fact collection.
+ \<open>success_intro\<close> fact collection.
@{const execute}, @{const effect}, @{const success} and @{const bind}
- are related by rules @{text execute_bind_success}, @{text
- success_bind_executeI}, @{text success_bind_effectI}, @{text
- effect_bindI}, @{text effect_bindE} and @{text execute_bind_eq_SomeI}.
-*}
+ are related by rules \<open>execute_bind_success\<close>, \<open>success_bind_executeI\<close>, \<open>success_bind_effectI\<close>, \<open>effect_bindI\<close>, \<open>effect_bindE\<close> and \<open>execute_bind_eq_SomeI\<close>.
+\<close>
-section {* Monadic data structures *}
+section \<open>Monadic data structures\<close>
-text {*
+text \<open>
The operations for monadic data structures (arrays and references)
come in two flavours:
@@ -136,11 +133,11 @@
Note that HOL equality coincides with reference equality and may be
used as primitive executable operation.
-*}
+\<close>
-subsection {* Arrays *}
+subsection \<open>Arrays\<close>
-text {*
+text \<open>
Heap operations:
\begin{quote}
@@ -166,11 +163,11 @@
@{term_type Array.swap} \\
@{term_type Array.freeze}
\end{quote}
-*}
+\<close>
-subsection {* References *}
+subsection \<open>References\<close>
-text {*
+text \<open>
Heap operations:
\begin{quote}
@@ -189,26 +186,26 @@
@{term_type Ref.update} \\
@{term_type Ref.change}
\end{quote}
-*}
+\<close>
-section {* Code generation *}
+section \<open>Code generation\<close>
-text {*
+text \<open>
Imperative HOL sets up the code generator in a way that imperative
operations are mapped to suitable counterparts in the target
- language. For @{text Haskell}, a suitable @{text ST} monad is used;
- for @{text SML}, @{text Ocaml} and @{text Scala} unit values ensure
+ language. For \<open>Haskell\<close>, a suitable \<open>ST\<close> monad is used;
+ for \<open>SML\<close>, \<open>Ocaml\<close> and \<open>Scala\<close> unit values ensure
that the evaluation order is the same as you would expect from the
original monadic expressions. These units may look cumbersome; the
- target language variants @{text SML_imp}, @{text Ocaml_imp} and
- @{text Scala_imp} make some effort to optimize some of them away.
-*}
+ target language variants \<open>SML_imp\<close>, \<open>Ocaml_imp\<close> and
+ \<open>Scala_imp\<close> make some effort to optimize some of them away.
+\<close>
-section {* Some hints for using the framework *}
+section \<open>Some hints for using the framework\<close>
-text {*
+text \<open>
Of course a framework itself does not by itself indicate how to make
best use of it. Here some hints drawn from prior experiences with
Imperative HOL:
@@ -228,8 +225,8 @@
consists of composing those.
\item Whether one should prefer equational reasoning (fact
- collection @{text execute_simps} or relational reasoning (fact
- collections @{text effect_intros} and @{text effect_elims}) depends
+ collection \<open>execute_simps\<close> or relational reasoning (fact
+ collections \<open>effect_intros\<close> and \<open>effect_elims\<close>) depends
on the problems to solve. For complex expressions or
expressions involving binders, the relation style usually is
superior but requires more proof text.
@@ -238,6 +235,6 @@
HOL yourself whenever appropriate.
\end{itemize}
-*}
+\<close>
end
--- a/src/HOL/Imperative_HOL/Ref.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Imperative_HOL/Ref.thy Thu May 26 17:51:22 2016 +0200
@@ -2,19 +2,19 @@
Author: John Matthews, Galois Connections; Alexander Krauss, Lukas Bulwahn & Florian Haftmann, TU Muenchen
*)
-section {* Monadic references *}
+section \<open>Monadic references\<close>
theory Ref
imports Array
begin
-text {*
+text \<open>
Imperative reference operations; modeled after their ML counterparts.
See @{url "http://caml.inria.fr/pub/docs/manual-caml-light/node14.15.html"}
and @{url "http://www.smlnj.org/doc/Conversion/top-level-comparison.html"}
-*}
+\<close>
-subsection {* Primitives *}
+subsection \<open>Primitives\<close>
definition present :: "heap \<Rightarrow> 'a::heap ref \<Rightarrow> bool" where
"present h r \<longleftrightarrow> addr_of_ref r < lim h"
@@ -36,7 +36,7 @@
"r =!= s \<longleftrightarrow> TYPEREP('a) \<noteq> TYPEREP('b) \<or> addr_of_ref r \<noteq> addr_of_ref s"
-subsection {* Monad operations *}
+subsection \<open>Monad operations\<close>
definition ref :: "'a::heap \<Rightarrow> 'a ref Heap" where
[code del]: "ref v = Heap_Monad.heap (alloc v)"
@@ -56,12 +56,12 @@
}"
-subsection {* Properties *}
+subsection \<open>Properties\<close>
-text {* Primitives *}
+text \<open>Primitives\<close>
lemma noteq_sym: "r =!= s \<Longrightarrow> s =!= r"
- and unequal [simp]: "r \<noteq> r' \<longleftrightarrow> r =!= r'" -- "same types!"
+ and unequal [simp]: "r \<noteq> r' \<longleftrightarrow> r =!= r'" \<comment> "same types!"
by (auto simp add: noteq_def)
lemma noteq_irrefl: "r =!= r \<Longrightarrow> False"
@@ -133,7 +133,7 @@
by (auto simp add: noteq_def present_def)
-text {* Monad operations *}
+text \<open>Monad operations\<close>
lemma execute_ref [execute_simps]:
"execute (ref v) h = Some (alloc v h)"
@@ -216,7 +216,7 @@
by (rule Heap_eqI) (simp add: change_def lookup_chain)
-text {* Non-interaction between imperative arrays and imperative references *}
+text \<open>Non-interaction between imperative arrays and imperative references\<close>
lemma array_get_set [simp]:
"Array.get (set r v h) = Array.get h"
@@ -261,9 +261,9 @@
hide_const (open) present get set alloc noteq lookup update change
-subsection {* Code generator setup *}
+subsection \<open>Code generator setup\<close>
-text {* Intermediate operation avoids invariance problem in @{text Scala} (similar to value restriction) *}
+text \<open>Intermediate operation avoids invariance problem in \<open>Scala\<close> (similar to value restriction)\<close>
definition ref' where
[code del]: "ref' = ref"
@@ -273,7 +273,7 @@
by (simp add: ref'_def)
-text {* SML / Eval *}
+text \<open>SML / Eval\<close>
code_printing type_constructor ref \<rightharpoonup> (SML) "_/ ref"
code_printing type_constructor ref \<rightharpoonup> (Eval) "_/ Unsynchronized.ref"
@@ -287,7 +287,7 @@
code_reserved Eval Unsynchronized
-text {* OCaml *}
+text \<open>OCaml\<close>
code_printing type_constructor ref \<rightharpoonup> (OCaml) "_/ ref"
code_printing constant Ref \<rightharpoonup> (OCaml) "failwith/ \"bare Ref\""
@@ -299,7 +299,7 @@
code_reserved OCaml ref
-text {* Haskell *}
+text \<open>Haskell\<close>
code_printing type_constructor ref \<rightharpoonup> (Haskell) "Heap.STRef/ Heap.RealWorld/ _"
code_printing constant Ref \<rightharpoonup> (Haskell) "error/ \"bare Ref\""
@@ -310,7 +310,7 @@
code_printing class_instance ref :: HOL.equal \<rightharpoonup> (Haskell) -
-text {* Scala *}
+text \<open>Scala\<close>
code_printing type_constructor ref \<rightharpoonup> (Scala) "!Ref[_]"
code_printing constant Ref \<rightharpoonup> (Scala) "!sys.error(\"bare Ref\")"
--- a/src/HOL/Imperative_HOL/ex/Imperative_Quicksort.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Imperative_HOL/ex/Imperative_Quicksort.thy Thu May 26 17:51:22 2016 +0200
@@ -2,7 +2,7 @@
Author: Lukas Bulwahn, TU Muenchen
*)
-section {* An imperative implementation of Quicksort on arrays *}
+section \<open>An imperative implementation of Quicksort on arrays\<close>
theory Imperative_Quicksort
imports
@@ -12,7 +12,7 @@
"~~/src/HOL/Library/Code_Target_Numeral"
begin
-text {* We prove QuickSort correct in the Relational Calculus. *}
+text \<open>We prove QuickSort correct in the Relational Calculus.\<close>
definition swap :: "'a::heap array \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> unit Heap"
where
@@ -76,11 +76,11 @@
using assms
proof (induct a l r p arbitrary: h h' rs rule:part1.induct)
case (1 a l r p h h' rs)
- note cr = `effect (part1 a l r p) h h' rs`
+ note cr = \<open>effect (part1 a l r p) h h' rs\<close>
show ?case
proof (cases "r \<le> l")
case True (* Terminating case *)
- with cr `l \<le> r` show ?thesis
+ with cr \<open>l \<le> r\<close> show ?thesis
unfolding part1.simps[of a l r p]
by (elim effect_bindE effect_ifE effect_returnE effect_nthE) auto
next
@@ -95,7 +95,7 @@
unfolding part1.simps[of a l r p]
by (elim effect_bindE effect_nthE effect_ifE effect_returnE) auto
from rec_condition have "l + 1 \<le> r" by arith
- from 1(1)[OF rec_condition True rec1 `l + 1 \<le> r`]
+ from 1(1)[OF rec_condition True rec1 \<open>l + 1 \<le> r\<close>]
show ?thesis by simp
next
case False
@@ -105,7 +105,7 @@
unfolding part1.simps[of a l r p]
by (elim effect_bindE effect_nthE effect_ifE effect_returnE) auto
from rec_condition have "l \<le> r - 1" by arith
- from 1(2) [OF rec_condition False rec2 `l \<le> r - 1`] show ?thesis by fastforce
+ from 1(2) [OF rec_condition False rec2 \<open>l \<le> r - 1\<close>] show ?thesis by fastforce
qed
qed
qed
@@ -116,7 +116,7 @@
using assms
proof (induct a l r p arbitrary: h h' rs rule:part1.induct)
case (1 a l r p h h' rs)
- note cr = `effect (part1 a l r p) h h' rs`
+ note cr = \<open>effect (part1 a l r p) h h' rs\<close>
show ?case
proof (cases "r \<le> l")
@@ -138,7 +138,7 @@
using assms
proof (induct a l r p arbitrary: h h' rs rule:part1.induct)
case (1 a l r p h h' rs)
- note cr = `effect (part1 a l r p) h h' rs`
+ note cr = \<open>effect (part1 a l r p) h h' rs\<close>
show ?case
proof (cases "r \<le> l")
@@ -183,7 +183,7 @@
using assms
proof (induct a l r p arbitrary: h h' rs rule:part1.induct)
case (1 a l r p h h' rs)
- note cr = `effect (part1 a l r p) h h' rs`
+ note cr = \<open>effect (part1 a l r p) h h' rs\<close>
show ?case
proof (cases "r \<le> l")
@@ -281,8 +281,8 @@
else middle)"
unfolding partition.simps
by (elim effect_bindE effect_returnE effect_nthE effect_ifE effect_updE) simp
- from `l < r` have "l \<le> r - 1" by arith
- from part_returns_index_in_bounds[OF part this] rs_equals `l < r` show ?thesis by auto
+ from \<open>l < r\<close> have "l \<le> r - 1" by arith
+ from part_returns_index_in_bounds[OF part this] rs_equals \<open>l < r\<close> show ?thesis by auto
qed
lemma partition_partitions:
@@ -307,9 +307,9 @@
by (elim effect_bindE effect_returnE effect_nthE effect_updE) simp
from swap have swap_length_remains: "Array.length h1 a = Array.length h' a"
unfolding swap_def by (elim effect_bindE effect_returnE effect_nthE effect_updE) auto
- from `l < r` have "l \<le> r - 1" by simp
+ from \<open>l < r\<close> have "l \<le> r - 1" by simp
note middle_in_bounds = part_returns_index_in_bounds[OF part this]
- from part_outer_remains[OF part] `l < r`
+ from part_outer_remains[OF part] \<open>l < r\<close>
have "Array.get h a ! r = Array.get h1 a ! r"
by fastforce
with swap
@@ -324,7 +324,7 @@
{
fix i
assume i_is_left: "l \<le> i \<and> i < rs"
- with swap_length_remains in_bounds middle_in_bounds rs_equals `l < r`
+ with swap_length_remains in_bounds middle_in_bounds rs_equals \<open>l < r\<close>
have i_props: "i < Array.length h' a" "i \<noteq> r" "i \<noteq> rs" by auto
from i_is_left rs_equals have "l \<le> i \<and> i < middle \<or> i = middle" by arith
with part_partitions[OF part] right_remains True
@@ -350,7 +350,7 @@
next
assume i_is: "rs < i \<and> i = r"
with rs_equals have "Suc middle \<noteq> r" by arith
- with middle_in_bounds `l < r` have "Suc middle \<le> r - 1" by arith
+ with middle_in_bounds \<open>l < r\<close> have "Suc middle \<le> r - 1" by arith
with part_partitions[OF part] right_remains
have "Array.get h' a ! rs \<le> Array.get h1 a ! (Suc middle)"
by fastforce
@@ -451,7 +451,7 @@
using assms
proof (induct a l r arbitrary: h h' rs rule: quicksort.induct)
case (1 a l r h h' rs)
- note cr = `effect (quicksort a l r) h h' rs`
+ note cr = \<open>effect (quicksort a l r) h h' rs\<close>
thus ?case
proof (cases "r > l")
case False
@@ -498,7 +498,7 @@
using assms
proof (induct a l r arbitrary: h h' rs rule: quicksort.induct)
case (1 a l r h h' rs)
- note cr = `effect (quicksort a l r) h h' rs`
+ note cr = \<open>effect (quicksort a l r) h h' rs\<close>
thus ?case
proof (cases "r > l")
case False
@@ -582,9 +582,9 @@
with quicksort_permutes [OF effect] properties_for_sort show ?thesis by fastforce
qed
-subsection {* No Errors in quicksort *}
-text {* We have proved that quicksort sorts (if no exceptions occur).
-We will now show that exceptions do not occur. *}
+subsection \<open>No Errors in quicksort\<close>
+text \<open>We have proved that quicksort sorts (if no exceptions occur).
+We will now show that exceptions do not occur.\<close>
lemma success_part1I:
assumes "l < Array.length h a" "r < Array.length h a"
@@ -647,7 +647,7 @@
qed
-subsection {* Example *}
+subsection \<open>Example\<close>
definition "qsort a = do {
k \<leftarrow> Array.len a;
@@ -662,7 +662,7 @@
qsort a
}"
-ML_val {* @{code example} () *}
+ML_val \<open>@{code example} ()\<close>
export_code qsort checking SML SML_imp OCaml? OCaml_imp? Haskell? Scala Scala_imp
--- a/src/HOL/Imperative_HOL/ex/Imperative_Reverse.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Imperative_HOL/ex/Imperative_Reverse.thy Thu May 26 17:51:22 2016 +0200
@@ -2,7 +2,7 @@
Author: Lukas Bulwahn, TU Muenchen
*)
-section {* An imperative in-place reversal on arrays *}
+section \<open>An imperative in-place reversal on arrays\<close>
theory Imperative_Reverse
imports Subarray "~~/src/HOL/Imperative_HOL/Imperative_HOL"
--- a/src/HOL/Imperative_HOL/ex/Linked_Lists.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Imperative_HOL/ex/Linked_Lists.thy Thu May 26 17:51:22 2016 +0200
@@ -2,15 +2,15 @@
Author: Lukas Bulwahn, TU Muenchen
*)
-section {* Linked Lists by ML references *}
+section \<open>Linked Lists by ML references\<close>
theory Linked_Lists
imports "../Imperative_HOL" "~~/src/HOL/Library/Code_Target_Int"
begin
-section {* Definition of Linked Lists *}
+section \<open>Definition of Linked Lists\<close>
-setup {* Sign.add_const_constraint (@{const_name Ref}, SOME @{typ "nat \<Rightarrow> 'a::type ref"}) *}
+setup \<open>Sign.add_const_constraint (@{const_name Ref}, SOME @{typ "nat \<Rightarrow> 'a::type ref"})\<close>
datatype 'a node = Empty | Node 'a "'a node ref"
primrec
@@ -55,9 +55,9 @@
by (simp_all add: traverse.simps[of "Empty"] traverse.simps[of "Node x r"])
-section {* Proving correctness with relational abstraction *}
+section \<open>Proving correctness with relational abstraction\<close>
-subsection {* Definition of list_of, list_of', refs_of and refs_of' *}
+subsection \<open>Definition of list_of, list_of', refs_of and refs_of'\<close>
primrec list_of :: "heap \<Rightarrow> ('a::heap) node \<Rightarrow> 'a list \<Rightarrow> bool"
where
@@ -79,7 +79,7 @@
| "refs_of' h r (x#xs) = ((x = r) \<and> refs_of h (Ref.get h x) xs)"
-subsection {* Properties of these definitions *}
+subsection \<open>Properties of these definitions\<close>
lemma list_of_Empty[simp]: "list_of h Empty xs = (xs = [])"
by (cases xs, auto)
@@ -216,7 +216,7 @@
thus ?thesis by (auto simp add: refs_of'_def')
qed
-subsection {* More complicated properties of these predicates *}
+subsection \<open>More complicated properties of these predicates\<close>
lemma list_of_append:
"list_of h n (as @ bs) \<Longrightarrow> \<exists>m. list_of h m bs"
@@ -258,7 +258,7 @@
by (fastforce simp add: refs_of_distinct refs_of_next)
-subsection {* Interaction of these predicates with our heap transitions *}
+subsection \<open>Interaction of these predicates with our heap transitions\<close>
lemma list_of_set_ref: "refs_of h q rs \<Longrightarrow> p \<notin> set rs \<Longrightarrow> list_of (Ref.set p v h) q as = list_of h q as"
proof (induct as arbitrary: q rs)
@@ -378,7 +378,7 @@
with assms that show thesis by auto
qed
-section {* Proving make_llist and traverse correct *}
+section \<open>Proving make_llist and traverse correct\<close>
lemma refs_of_invariant:
assumes "refs_of h (r::('a::heap) node) xs"
@@ -518,9 +518,9 @@
using effect_deterministic by fastforce
qed
-section {* Proving correctness of in-place reversal *}
+section \<open>Proving correctness of in-place reversal\<close>
-subsection {* Definition of in-place reversal *}
+subsection \<open>Definition of in-place reversal\<close>
partial_function (heap) rev' :: "('a::heap) node ref \<Rightarrow> 'a node ref \<Rightarrow> 'a node ref Heap"
where
@@ -541,7 +541,7 @@
"rev Empty = return Empty"
| "rev (Node x n) = do { q \<leftarrow> ref Empty; p \<leftarrow> ref (Node x n); v \<leftarrow> rev' q p; !v }"
-subsection {* Correctness Proof *}
+subsection \<open>Correctness Proof\<close>
lemma rev'_invariant:
assumes "effect (rev' q p) h h' v"
@@ -649,10 +649,10 @@
qed
-section {* The merge function on Linked Lists *}
-text {* We also prove merge correct *}
+section \<open>The merge function on Linked Lists\<close>
+text \<open>We also prove merge correct\<close>
-text{* First, we define merge on lists in a natural way. *}
+text\<open>First, we define merge on lists in a natural way.\<close>
fun Lmerge :: "('a::ord) list \<Rightarrow> 'a list \<Rightarrow> 'a list"
where
@@ -661,7 +661,7 @@
| "Lmerge [] ys = ys"
| "Lmerge xs [] = xs"
-subsection {* Definition of merge function *}
+subsection \<open>Definition of merge function\<close>
partial_function (heap) merge :: "('a::{heap, ord}) node ref \<Rightarrow> 'a node ref \<Rightarrow> 'a node ref Heap"
where
@@ -694,9 +694,9 @@
lemma sum_distrib: "case_sum fl fr (case x of Empty \<Rightarrow> y | Node v n \<Rightarrow> (z v n)) = (case x of Empty \<Rightarrow> case_sum fl fr y | Node v n \<Rightarrow> case_sum fl fr (z v n))"
by (cases x) auto
-subsection {* Induction refinement by applying the abstraction function to our induct rule *}
+subsection \<open>Induction refinement by applying the abstraction function to our induct rule\<close>
-text {* From our original induction rule Lmerge.induct, we derive a new rule with our list_of' predicate *}
+text \<open>From our original induction rule Lmerge.induct, we derive a new rule with our list_of' predicate\<close>
lemma merge_induct2:
assumes "list_of' h (p::'a::{heap, ord} node ref) xs"
@@ -741,7 +741,7 @@
qed
-text {* secondly, we add the effect statement in the premise, and derive the effect statements for the single cases which we then eliminate with our effect elim rules. *}
+text \<open>secondly, we add the effect statement in the premise, and derive the effect statements for the single cases which we then eliminate with our effect elim rules.\<close>
lemma merge_induct3:
assumes "list_of' h p xs"
@@ -790,10 +790,10 @@
qed
-subsection {* Proving merge correct *}
+subsection \<open>Proving merge correct\<close>
-text {* As many parts of the following three proofs are identical, we could actually move the
-same reasoning into an extended induction rule *}
+text \<open>As many parts of the following three proofs are identical, we could actually move the
+same reasoning into an extended induction rule\<close>
lemma merge_unchanged:
assumes "refs_of' h p xs"
@@ -822,7 +822,7 @@
from 3(11) pnrs_def have no_inter: "set pnrs \<inter> set ys = {}" by auto
from 3(7)[OF refs_of'_pn 3(10) this p_in] 3(3) have p_is_Node: "Ref.get h1 p = Node x pn"
by simp
- from 3(7)[OF refs_of'_pn 3(10) no_inter r_in] 3(8) `r \<noteq> p` show ?case
+ from 3(7)[OF refs_of'_pn 3(10) no_inter r_in] 3(8) \<open>r \<noteq> p\<close> show ?case
by simp
next
case (4 x xs' y ys' p q pn qn h1 r1 h' xs ys r)
@@ -835,7 +835,7 @@
with 4(11) 4(12) qnrs_def refs_of'_distinct[OF 4(10)] have q_in: "q \<notin> set xs \<union> set qnrs" by auto
from 4(11) qnrs_def have no_inter: "set xs \<inter> set qnrs = {}" by auto
from 4(7)[OF 4(9) refs_of'_qn this q_in] 4(4) have q_is_Node: "Ref.get h1 q = Node y qn" by simp
- from 4(7)[OF 4(9) refs_of'_qn no_inter r_in] 4(8) `r \<noteq> q` show ?case
+ from 4(7)[OF 4(9) refs_of'_qn no_inter r_in] 4(8) \<open>r \<noteq> q\<close> show ?case
by simp
qed
qed
@@ -937,9 +937,9 @@
show ?case by (auto simp: list_of'_set_ref)
qed
-section {* Code generation *}
+section \<open>Code generation\<close>
-text {* A simple example program *}
+text \<open>A simple example program\<close>
definition test_1 where "test_1 = (do { ll_xs \<leftarrow> make_llist [1..(15::int)]; xs \<leftarrow> traverse ll_xs; return xs })"
definition test_2 where "test_2 = (do { ll_xs \<leftarrow> make_llist [1..(15::int)]; ll_ys \<leftarrow> rev ll_xs; ys \<leftarrow> traverse ll_ys; return ys })"
@@ -958,9 +958,9 @@
code_reserved SML upto
-ML_val {* @{code test_1} () *}
-ML_val {* @{code test_2} () *}
-ML_val {* @{code test_3} () *}
+ML_val \<open>@{code test_1} ()\<close>
+ML_val \<open>@{code test_2} ()\<close>
+ML_val \<open>@{code test_3} ()\<close>
export_code test_1 test_2 test_3 checking SML SML_imp OCaml? OCaml_imp? Haskell? Scala Scala_imp
--- a/src/HOL/Imperative_HOL/ex/List_Sublist.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Imperative_HOL/ex/List_Sublist.thy Thu May 26 17:51:22 2016 +0200
@@ -2,7 +2,7 @@
Author: Lukas Bulwahn, TU Muenchen
*)
-section {* Slices of lists *}
+section \<open>Slices of lists\<close>
theory List_Sublist
imports "~~/src/HOL/Library/Multiset"
@@ -188,7 +188,7 @@
apply auto
done
-section {* Another sublist function *}
+section \<open>Another sublist function\<close>
function sublist' :: "nat \<Rightarrow> nat \<Rightarrow> 'a list \<Rightarrow> 'a list"
where
@@ -199,7 +199,7 @@
by pat_completeness auto
termination by lexicographic_order
-subsection {* Proving equivalence to the other sublist command *}
+subsection \<open>Proving equivalence to the other sublist command\<close>
lemma sublist'_sublist: "sublist' n m xs = sublist xs {j. n \<le> j \<and> j < m}"
apply (induct xs arbitrary: n m)
@@ -227,7 +227,7 @@
done
-subsection {* Showing equivalence to use of drop and take for definition *}
+subsection \<open>Showing equivalence to use of drop and take for definition\<close>
lemma "sublist' n m xs = take (m - n) (drop n xs)"
apply (induct xs arbitrary: n m)
@@ -239,7 +239,7 @@
apply simp
done
-subsection {* General lemma about sublist *}
+subsection \<open>General lemma about sublist\<close>
lemma sublist'_Nil[simp]: "sublist' i j [] = []"
by simp
--- a/src/HOL/Imperative_HOL/ex/SatChecker.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Imperative_HOL/ex/SatChecker.thy Thu May 26 17:51:22 2016 +0200
@@ -2,22 +2,22 @@
Author: Lukas Bulwahn, TU Muenchen
*)
-section {* An efficient checker for proofs from a SAT solver *}
+section \<open>An efficient checker for proofs from a SAT solver\<close>
theory SatChecker
imports "~~/src/HOL/Library/RBT_Impl" Sorted_List "../Imperative_HOL"
begin
-section{* General settings and functions for our representation of clauses *}
+section\<open>General settings and functions for our representation of clauses\<close>
-subsection{* Types for Literals, Clauses and ProofSteps *}
-text {* We encode Literals as integers and Clauses as sorted Lists. *}
+subsection\<open>Types for Literals, Clauses and ProofSteps\<close>
+text \<open>We encode Literals as integers and Clauses as sorted Lists.\<close>
type_synonym ClauseId = nat
type_synonym Lit = int
type_synonym Clause = "Lit list"
-text {* This resembles exactly to Isat's Proof Steps *}
+text \<open>This resembles exactly to Isat's Proof Steps\<close>
type_synonym Resolvants = "ClauseId * (Lit * ClauseId) list"
datatype ProofStep =
@@ -27,8 +27,8 @@
| Delete ClauseId
| Xstep ClauseId ClauseId
-subsection{* Interpretation of Literals, Clauses, and an array of Clauses *}
-text {* Specific definitions for Literals as integers *}
+subsection\<open>Interpretation of Literals, Clauses, and an array of Clauses\<close>
+text \<open>Specific definitions for Literals as integers\<close>
definition compl :: "Lit \<Rightarrow> Lit"
where
@@ -54,7 +54,7 @@
lemma compl_exists: "\<exists>l'. l = compl l'"
unfolding compl_def by arith
-text {* Specific definitions for Clauses as sorted lists *}
+text \<open>Specific definitions for Clauses as sorted lists\<close>
definition interpClause :: "(nat \<Rightarrow> bool) \<Rightarrow> Clause \<Rightarrow> bool"
where
@@ -116,12 +116,12 @@
unfolding inconsistent_def correctClause_def
by auto
-text {* Specific definition for derived clauses in the Array *}
+text \<open>Specific definition for derived clauses in the Array\<close>
definition
array_ran :: "('a::heap) option array \<Rightarrow> heap \<Rightarrow> 'a set" where
"array_ran a h = {e. Some e \<in> set (Array.get h a)}"
- -- {*FIXME*}
+ \<comment> \<open>FIXME\<close>
lemma array_ranI: "\<lbrakk> Some b = Array.get h a ! i; i < Array.length h a \<rbrakk> \<Longrightarrow> b \<in> array_ran a h"
unfolding array_ran_def Array.length_def by simp
@@ -165,11 +165,11 @@
by auto
-section{* Improved version of SatChecker *}
+section\<open>Improved version of SatChecker\<close>
-text{* This version just uses a single list traversal. *}
+text\<open>This version just uses a single list traversal.\<close>
-subsection{* Function definitions *}
+subsection\<open>Function definitions\<close>
primrec res_mem :: "Lit \<Rightarrow> Clause \<Rightarrow> Clause Heap"
where
@@ -209,7 +209,7 @@
declare res_mem.simps [simp del] resolve1.simps [simp del] resolve2.simps [simp del] res_thm'.simps [simp del]
-subsection {* Proofs about these functions *}
+subsection \<open>Proofs about these functions\<close>
lemma res_mem:
assumes "effect (res_mem l xs) h h' r"
@@ -408,7 +408,7 @@
done
qed
-subsection {* res_thm and doProofStep *}
+subsection \<open>res_thm and doProofStep\<close>
definition get_clause :: "Clause option array \<Rightarrow> ClauseId \<Rightarrow> Clause Heap"
where
@@ -598,9 +598,9 @@
apply (drule bspec)
by (rule array_ranI, auto)
-section {* Functional version with Lists as Array *}
+section \<open>Functional version with Lists as Array\<close>
-subsection {* List specific definitions *}
+subsection \<open>List specific definitions\<close>
definition list_ran :: "'a option list \<Rightarrow> 'a set"
where
@@ -636,7 +636,7 @@
"correctList rootcls xs =
(\<forall>cl \<in> list_ran xs. correctClause rootcls cl \<and> sorted cl \<and> distinct cl)"
-subsection {* Checker functions *}
+subsection \<open>Checker functions\<close>
primrec lres_thm :: "Clause option list \<Rightarrow> (Lit * ClauseId) \<Rightarrow> Clause \<Rightarrow> Clause Heap"
where
@@ -670,7 +670,7 @@
}"
-section {* Functional version with RedBlackTrees *}
+section \<open>Functional version with RedBlackTrees\<close>
primrec tres_thm :: "(ClauseId, Clause) RBT_Impl.rbt \<Rightarrow> Lit \<times> ClauseId \<Rightarrow> Clause \<Rightarrow> Clause Heap"
where
--- a/src/HOL/Imperative_HOL/ex/Sorted_List.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Imperative_HOL/ex/Sorted_List.thy Thu May 26 17:51:22 2016 +0200
@@ -2,13 +2,13 @@
Author: Lukas Bulwahn, TU Muenchen
*)
-section {* Sorted lists as representation of finite sets *}
+section \<open>Sorted lists as representation of finite sets\<close>
theory Sorted_List
imports Main
begin
-text {* Merge function for two distinct sorted lists to get compound distinct sorted list *}
+text \<open>Merge function for two distinct sorted lists to get compound distinct sorted list\<close>
fun merge :: "('a::linorder) list \<Rightarrow> 'a list \<Rightarrow> 'a list"
where
@@ -17,7 +17,7 @@
| "merge xs [] = xs"
| "merge [] ys = ys"
-text {* The function package does not derive automatically the more general rewrite rule as follows: *}
+text \<open>The function package does not derive automatically the more general rewrite rule as follows:\<close>
lemma merge_Nil[simp]: "merge [] ys = ys"
by (cases ys) auto
@@ -31,7 +31,7 @@
lemma distinct_merge[simp]: "\<lbrakk> distinct xs; distinct ys; List.sorted xs; List.sorted ys \<rbrakk> \<Longrightarrow> distinct (merge xs ys)"
by (induct xs ys rule: merge.induct, auto simp add: sorted_Cons)
-text {* The remove function removes an element from a sorted list *}
+text \<open>The remove function removes an element from a sorted list\<close>
primrec remove :: "('a :: linorder) \<Rightarrow> 'a list \<Rightarrow> 'a list"
where
@@ -86,7 +86,7 @@
apply (auto simp add: sorted_Cons)
done
-subsection {* Efficient member function for sorted lists *}
+subsection \<open>Efficient member function for sorted lists\<close>
primrec smember :: "'a list \<Rightarrow> 'a::linorder \<Rightarrow> bool" where
"smember [] x \<longleftrightarrow> False"
--- a/src/HOL/Imperative_HOL/ex/Subarray.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Imperative_HOL/ex/Subarray.thy Thu May 26 17:51:22 2016 +0200
@@ -2,7 +2,7 @@
Author: Lukas Bulwahn, TU Muenchen
*)
-section {* Theorems about sub arrays *}
+section \<open>Theorems about sub arrays\<close>
theory Subarray
imports "~~/src/HOL/Imperative_HOL/Array" List_Sublist
--- a/src/HOL/Import/HOL_Light_Import.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Import/HOL_Light_Import.thy Thu May 26 17:51:22 2016 +0200
@@ -3,7 +3,7 @@
Author: Alexander Krauss, QAware GmbH
*)
-section {* Main HOL Light importer *}
+section \<open>Main HOL Light importer\<close>
theory HOL_Light_Import
imports HOL_Light_Maps
--- a/src/HOL/Import/HOL_Light_Maps.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Import/HOL_Light_Maps.thy Thu May 26 17:51:22 2016 +0200
@@ -5,7 +5,7 @@
Based on earlier code by Steven Obua and Sebastian Skalberg
*)
-section {* Type and constant mappings of HOL Light importer *}
+section \<open>Type and constant mappings of HOL Light importer\<close>
theory HOL_Light_Maps
imports Import_Setup
--- a/src/HOL/Import/Import_Setup.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Import/Import_Setup.thy Thu May 26 17:51:22 2016 +0200
@@ -3,7 +3,7 @@
Author: Alexander Krauss, QAware GmbH
*)
-section {* Importer machinery and required theorems *}
+section \<open>Importer machinery and required theorems\<close>
theory Import_Setup
imports Main "~~/src/HOL/Binomial"
--- a/src/HOL/Induct/ABexp.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Induct/ABexp.thy Thu May 26 17:51:22 2016 +0200
@@ -55,7 +55,7 @@
"evala env (substa (Var (v := a')) a) = evala (env (v := evala env a')) a"
and subst1_bexp:
"evalb env (substb (Var (v := a')) b) = evalb (env (v := evala env a')) b"
- -- \<open>one variable\<close>
+ \<comment> \<open>one variable\<close>
by (induct a and b) simp_all
lemma subst_all_aexp:
--- a/src/HOL/Induct/Com.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Induct/Com.thy Thu May 26 17:51:22 2016 +0200
@@ -136,7 +136,7 @@
by (rule fun_upd_same [THEN subst]) fast
-text\<open>Make the induction rule look nicer -- though @{text eta_contract} makes the new
+text\<open>Make the induction rule look nicer -- though \<open>eta_contract\<close> makes the new
version look worse than it is...\<close>
lemma split_lemma: "{((e,s),(n,s')). P e s n s'} = Collect (case_prod (%v. case_prod (case_prod P v)))"
@@ -167,11 +167,11 @@
done
-text\<open>Lemma for @{text Function_eval}. The major premise is that @{text "(c,s)"} executes to @{text "s1"}
+text\<open>Lemma for \<open>Function_eval\<close>. The major premise is that \<open>(c,s)\<close> executes to \<open>s1\<close>
using eval restricted to its functional part. Note that the execution
- @{text "(c,s) -[eval]-> s2"} can use unrestricted @{text eval}! The reason is that
- the execution @{text "(c,s) -[eval Int {...}]-> s1"} assures us that execution is
- functional on the argument @{text "(c,s)"}.
+ \<open>(c,s) -[eval]-> s2\<close> can use unrestricted \<open>eval\<close>! The reason is that
+ the execution \<open>(c,s) -[eval Int {...}]-> s1\<close> assures us that execution is
+ functional on the argument \<open>(c,s)\<close>.
\<close>
lemma com_Unique:
"(c,s) -[eval Int {((e,s),(n,t)). \<forall>nt'. (e,s) -|-> nt' --> (n,t)=nt'}]-> s1
--- a/src/HOL/Induct/Common_Patterns.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Induct/Common_Patterns.thy Thu May 26 17:51:22 2016 +0200
@@ -10,7 +10,7 @@
text \<open>
The subsequent Isar proof schemes illustrate common proof patterns
- supported by the generic @{text "induct"} method.
+ supported by the generic \<open>induct\<close> method.
To demonstrate variations on statement (goal) structure we refer to
the induction rule of Peano natural numbers: @{thm nat.induct
@@ -29,8 +29,8 @@
text \<open>
Augmenting a problem by additional facts and locally fixed variables
is a bread-and-butter method in many applications. This is where
- unwieldy object-level @{text "\<forall>"} and @{text "\<longrightarrow>"} used to occur in
- the past. The @{text "induct"} method works with primary means of
+ unwieldy object-level \<open>\<forall>\<close> and \<open>\<longrightarrow>\<close> used to occur in
+ the past. The \<open>induct\<close> method works with primary means of
the proof language instead.
\<close>
@@ -80,8 +80,8 @@
qed
text \<open>
- Observe how the local definition @{text "n = a x"} recurs in the
- inductive cases as @{text "0 = a x"} and @{text "Suc n = a x"},
+ Observe how the local definition \<open>n = a x\<close> recurs in the
+ inductive cases as \<open>0 = a x\<close> and \<open>Suc n = a x\<close>,
according to underlying induction rule.
\<close>
@@ -155,9 +155,9 @@
The following pattern illustrates the slightly more complex
situation of simultaneous goals with individual local assumptions.
In compound simultaneous statements like this, local assumptions
- need to be included into each goal, using @{text "\<Longrightarrow>"} of the Pure
+ need to be included into each goal, using \<open>\<Longrightarrow>\<close> of the Pure
framework. In contrast, local parameters do not require separate
- @{text "\<And>"} prefixes here, but may be moved into the common context
+ \<open>\<And>\<close> prefixes here, but may be moved into the common context
of the whole statement.
\<close>
@@ -196,7 +196,7 @@
qed
text \<open>
- Here @{text "induct"} provides again nested cases with numbered
+ Here \<open>induct\<close> provides again nested cases with numbered
sub-cases, which allows to share common parts of the body context.
In typical applications, there could be a long intermediate proof of
general consequences of the induction hypotheses, before finishing
@@ -208,8 +208,7 @@
text \<open>
Multiple induction rules emerge from mutual definitions of
- datatypes, inductive predicates, functions etc. The @{text
- "induct"} method accepts replicated arguments (with @{text "and"}
+ datatypes, inductive predicates, functions etc. The \<open>induct\<close> method accepts replicated arguments (with \<open>and\<close>
separator), corresponding to each projection of the induction
principle.
@@ -344,7 +343,7 @@
text \<open>
Working with mutual rules requires special care in composing the
statement as a two-level conjunction, using lists of propositions
- separated by @{text "and"}. For example:
+ separated by \<open>and\<close>. For example:
\<close>
inductive Evn :: "nat \<Rightarrow> bool" and Odd :: "nat \<Rightarrow> bool"
--- a/src/HOL/Induct/PropLog.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Induct/PropLog.thy Thu May 26 17:51:22 2016 +0200
@@ -14,8 +14,8 @@
Inductive definition of propositional logic. Soundness and
completeness w.r.t.\ truth-tables.
- Prove: If @{text "H |= p"} then @{text "G |= p"} where @{text "G \<in>
- Fin(H)"}
+ Prove: If \<open>H |= p\<close> then \<open>G |= p\<close> where \<open>G \<in>
+ Fin(H)\<close>
\<close>
subsection \<open>The datatype of propositions\<close>
@@ -49,8 +49,8 @@
| eval_imp: "tt[[p->q]] = (tt[[p]] --> tt[[q]])"
text \<open>
- A finite set of hypotheses from @{text t} and the @{text Var}s in
- @{text p}.
+ A finite set of hypotheses from \<open>t\<close> and the \<open>Var\<close>s in
+ \<open>p\<close>.
\<close>
primrec hyps :: "['a pl, 'a set] => 'a pl set"
@@ -63,8 +63,8 @@
subsubsection \<open>Logical consequence\<close>
text \<open>
- For every valuation, if all elements of @{text H} are true then so
- is @{text p}.
+ For every valuation, if all elements of \<open>H\<close> are true then so
+ is \<open>p\<close>.
\<close>
definition sat :: "['a pl set, 'a pl] => bool" (infixl "|=" 50)
@@ -80,14 +80,14 @@
done
lemma thms_I: "H |- p->p"
- -- \<open>Called @{text I} for Identity Combinator, not for Introduction.\<close>
+ \<comment> \<open>Called \<open>I\<close> for Identity Combinator, not for Introduction.\<close>
by (best intro: thms.K thms.S thms.MP)
subsubsection \<open>Weakening, left and right\<close>
lemma weaken_left: "[| G \<subseteq> H; G|-p |] ==> H|-p"
- -- \<open>Order of premises is convenient with @{text THEN}\<close>
+ \<comment> \<open>Order of premises is convenient with \<open>THEN\<close>\<close>
by (erule thms_mono [THEN predicate1D])
lemma weaken_left_insert: "G |- p \<Longrightarrow> insert a G |- p"
@@ -146,7 +146,7 @@
done
lemma hyps_thms_if: "hyps p tt |- (if tt[[p]] then p else p->false)"
- -- \<open>Typical example of strengthening the induction statement.\<close>
+ \<comment> \<open>Typical example of strengthening the induction statement.\<close>
apply simp
apply (induct p)
apply (simp_all add: thms_I thms.H)
@@ -155,8 +155,8 @@
done
lemma sat_thms_p: "{} |= p ==> hyps p tt |- p"
- -- \<open>Key lemma for completeness; yields a set of assumptions
- satisfying @{text p}\<close>
+ \<comment> \<open>Key lemma for completeness; yields a set of assumptions
+ satisfying \<open>p\<close>\<close>
unfolding sat_def
by (metis (full_types) empty_iff hyps_thms_if)
@@ -178,7 +178,7 @@
lemma thms_excluded_middle_rule:
"[| insert p H |- q; insert (p->false) H |- q |] ==> H |- q"
- -- \<open>Hard to prove directly because it requires cuts\<close>
+ \<comment> \<open>Hard to prove directly because it requires cuts\<close>
by (rule thms_excluded_middle [THEN thms.MP, THEN thms.MP], auto)
@@ -200,7 +200,7 @@
lemma hyps_insert: "hyps p (insert v t) <= insert (#v) (hyps p t-{#v->false})"
by (induct p) auto
-text \<open>Two lemmas for use with @{text weaken_left}\<close>
+text \<open>Two lemmas for use with \<open>weaken_left\<close>\<close>
lemma insert_Diff_same: "B-C <= insert a (B-insert a C)"
by fast
@@ -234,12 +234,12 @@
"{} |= p ==> \<forall>t. hyps p t - hyps p t0 |- p"
apply (rule hyps_subset [THEN hyps_finite [THEN finite_subset_induct]])
apply (simp add: sat_thms_p, safe)
- txt\<open>Case @{text"hyps p t-insert(#v,Y) |- p"}\<close>
+ txt\<open>Case \<open>hyps p t-insert(#v,Y) |- p\<close>\<close>
apply (iprover intro: thms_excluded_middle_rule
insert_Diff_same [THEN weaken_left]
insert_Diff_subset2 [THEN weaken_left]
hyps_Diff [THEN Diff_weaken_left])
-txt\<open>Case @{text"hyps p t-insert(#v -> false,Y) |- p"}\<close>
+txt\<open>Case \<open>hyps p t-insert(#v -> false,Y) |- p\<close>\<close>
apply (iprover intro: thms_excluded_middle_rule
insert_Diff_same [THEN weaken_left]
insert_Diff_subset2 [THEN weaken_left]
--- a/src/HOL/Induct/QuoDataType.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Induct/QuoDataType.thy Thu May 26 17:51:22 2016 +0200
@@ -420,7 +420,7 @@
subsection\<open>The Abstract Discriminator\<close>
-text\<open>However, as @{text Crypt_Nonce_neq_Nonce} above illustrates, we don't
+text\<open>However, as \<open>Crypt_Nonce_neq_Nonce\<close> above illustrates, we don't
need this function in order to prove discrimination theorems.\<close>
definition
--- a/src/HOL/Induct/QuoNestedDataType.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Induct/QuoNestedDataType.thy Thu May 26 17:51:22 2016 +0200
@@ -241,7 +241,7 @@
regard an @{typ "exp list"} as a @{term "listrel exprel"} equivalence class\<close>
text\<open>This theorem is easily proved but never used. There's no obvious way
-even to state the analogous result, @{text FnCall_Cons}.\<close>
+even to state the analogous result, \<open>FnCall_Cons\<close>.\<close>
lemma FnCall_Nil: "FnCall F [] = Abs_Exp (exprel``{FNCALL F []})"
by (simp add: FnCall_def)
@@ -390,7 +390,7 @@
subsection\<open>The Abstract Discriminator\<close>
-text\<open>However, as @{text FnCall_Var_neq_Var} illustrates, we don't need this
+text\<open>However, as \<open>FnCall_Var_neq_Var\<close> illustrates, we don't need this
function in order to prove discrimination theorems.\<close>
definition
--- a/src/HOL/Induct/Sigma_Algebra.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Induct/Sigma_Algebra.thy Thu May 26 17:51:22 2016 +0200
@@ -10,8 +10,7 @@
text \<open>
This is just a tiny example demonstrating the use of inductive
- definitions in classical mathematics. We define the least @{text
- \<sigma>}-algebra over a given set of sets.
+ definitions in classical mathematics. We define the least \<open>\<sigma>\<close>-algebra over a given set of sets.
\<close>
inductive_set \<sigma>_algebra :: "'a set set \<Rightarrow> 'a set set"
@@ -24,7 +23,7 @@
text \<open>
The following basic facts are consequences of the closure properties
- of any @{text \<sigma>}-algebra, merely using the introduction rules, but
+ of any \<open>\<sigma>\<close>-algebra, merely using the introduction rules, but
no induction nor cases.
\<close>
--- a/src/HOL/Matrix_LP/ComputeFloat.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Matrix_LP/ComputeFloat.thy Thu May 26 17:51:22 2016 +0200
@@ -2,7 +2,7 @@
Author: Steven Obua
*)
-section {* Floating Point Representation of the Reals *}
+section \<open>Floating Point Representation of the Reals\<close>
theory ComputeFloat
imports Complex_Main "~~/src/HOL/Library/Lattice_Algebras"
--- a/src/HOL/Matrix_LP/ComputeHOL.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Matrix_LP/ComputeHOL.thy Thu May 26 17:51:22 2016 +0200
@@ -134,7 +134,7 @@
lemmas compute_hol = compute_if compute_bool compute_pair compute_option compute_list compute_let
-ML {*
+ML \<open>
signature ComputeHOL =
sig
val prep_thms : thm list -> thm list
@@ -182,6 +182,6 @@
end
end
-*}
+\<close>
end
--- a/src/HOL/Matrix_LP/Matrix.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Matrix_LP/Matrix.thy Thu May 26 17:51:22 2016 +0200
@@ -265,13 +265,13 @@
definition associative :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> bool" where
"associative f == ! x y z. f (f x y) z = f x (f y z)"
-text{*
+text\<open>
To reason about associativity and commutativity of operations on matrices,
let's take a step back and look at the general situtation: Assume that we have
sets $A$ and $B$ with $B \subset A$ and an abstraction $u: A \rightarrow B$. This abstraction has to fulfill $u(b) = b$ for all $b \in B$, but is arbitrary otherwise.
Each function $f: A \times A \rightarrow A$ now induces a function $f': B \times B \rightarrow B$ by $f' = u \circ f$.
It is obvious that commutativity of $f$ implies commutativity of $f'$: $f' x y = u (f x y) = u (f y x) = f' y x.$
-*}
+\<close>
lemma combine_infmatrix_commute:
"commutative f \<Longrightarrow> commutative (combine_infmatrix f)"
@@ -281,14 +281,14 @@
"commutative f \<Longrightarrow> commutative (combine_matrix f)"
by (simp add: combine_matrix_def commutative_def combine_infmatrix_def)
-text{*
+text\<open>
On the contrary, given an associative function $f$ we cannot expect $f'$ to be associative. A counterexample is given by $A=\ganz$, $B=\{-1, 0, 1\}$,
as $f$ we take addition on $\ganz$, which is clearly associative. The abstraction is given by $u(a) = 0$ for $a \notin B$. Then we have
\[ f' (f' 1 1) -1 = u(f (u (f 1 1)) -1) = u(f (u 2) -1) = u (f 0 -1) = -1, \]
but on the other hand we have
\[ f' 1 (f' 1 -1) = u (f 1 (u (f 1 -1))) = u (f 1 0) = 1.\]
A way out of this problem is to assume that $f(A\times A)\subset A$ holds, and this is what we are going to do:
-*}
+\<close>
lemma nonzero_positions_combine_infmatrix[simp]: "f 0 0 = 0 \<Longrightarrow> nonzero_positions (combine_infmatrix f A B) \<subseteq> (nonzero_positions A) \<union> (nonzero_positions B)"
by (rule subsetI, simp add: nonzero_positions_def combine_infmatrix_def, auto)
@@ -303,7 +303,7 @@
apply (rule finite_subset[of _ "(nonzero_positions (Rep_matrix A)) \<union> (nonzero_positions (Rep_matrix B))"])
by (simp_all)
-text {* We need the next two lemmas only later, but it is analog to the above one, so we prove them now: *}
+text \<open>We need the next two lemmas only later, but it is analog to the above one, so we prove them now:\<close>
lemma nonzero_positions_apply_infmatrix[simp]: "f 0 = 0 \<Longrightarrow> nonzero_positions (apply_infmatrix f A) \<subseteq> nonzero_positions A"
by (rule subsetI, simp add: nonzero_positions_def apply_infmatrix_def, auto)
--- a/src/HOL/Metis_Examples/Abstraction.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Metis_Examples/Abstraction.thy Thu May 26 17:51:22 2016 +0200
@@ -5,7 +5,7 @@
Example featuring Metis's support for lambda-abstractions.
*)
-section {* Example Featuring Metis's Support for Lambda-Abstractions *}
+section \<open>Example Featuring Metis's Support for Lambda-Abstractions\<close>
theory Abstraction
imports "~~/src/HOL/Library/FuncSet"
--- a/src/HOL/Metis_Examples/Big_O.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Metis_Examples/Big_O.thy Thu May 26 17:51:22 2016 +0200
@@ -5,7 +5,7 @@
Metis example featuring the Big O notation.
*)
-section {* Metis Example Featuring the Big O Notation *}
+section \<open>Metis Example Featuring the Big O Notation\<close>
theory Big_O
imports
@@ -14,7 +14,7 @@
"~~/src/HOL/Library/Set_Algebras"
begin
-subsection {* Definitions *}
+subsection \<open>Definitions\<close>
definition bigo :: "('a => 'b::linordered_idom) => ('a => 'b) set" ("(1O'(_'))") where
"O(f::('a => 'b)) == {h. \<exists>c. \<forall>x. \<bar>h x\<bar> <= c * \<bar>f x\<bar>}"
@@ -457,7 +457,7 @@
hence "\<exists>(v::'a) (u::'a) SKF\<^sub>7::'a \<Rightarrow> 'b. \<bar>inverse c\<bar> * \<bar>g (SKF\<^sub>7 (u * v))\<bar> \<le> u * (v * \<bar>f (SKF\<^sub>7 (u * v))\<bar>)"
by (metis mult_left_mono)
then show "\<exists>ca::'a. \<forall>x::'b. inverse \<bar>c\<bar> * \<bar>g x\<bar> \<le> ca * \<bar>f x\<bar>"
- using A2 F4 by (metis F1 `0 < \<bar>inverse c\<bar>` linordered_field_class.sign_simps(23) mult_le_cancel_left_pos)
+ using A2 F4 by (metis F1 \<open>0 < \<bar>inverse c\<bar>\<close> linordered_field_class.sign_simps(23) mult_le_cancel_left_pos)
qed
lemma bigo_const_mult6 [intro]: "(\<lambda>x. c) *o O(f) <= O(f)"
@@ -489,7 +489,7 @@
apply (simp add: fun_diff_def)
done
-subsection {* Setsum *}
+subsection \<open>Setsum\<close>
lemma bigo_setsum_main: "\<forall>x. \<forall>y \<in> A x. 0 <= h x y \<Longrightarrow>
\<exists>c. \<forall>x. \<forall>y \<in> A x. \<bar>f x y\<bar> <= c * (h x y) \<Longrightarrow>
@@ -566,7 +566,7 @@
apply auto
done
-subsection {* Misc useful stuff *}
+subsection \<open>Misc useful stuff\<close>
lemma bigo_useful_intro: "A <= O(f) \<Longrightarrow> B <= O(f) \<Longrightarrow>
A + B <= O(f)"
@@ -610,7 +610,7 @@
apply (simp add: fun_diff_def)
done
-subsection {* Less than or equal to *}
+subsection \<open>Less than or equal to\<close>
definition lesso :: "('a => 'b::linordered_idom) => ('a => 'b) => ('a => 'b)" (infixl "<o" 70) where
"f <o g == (\<lambda>x. max (f x - g x) 0)"
--- a/src/HOL/Metis_Examples/Binary_Tree.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Metis_Examples/Binary_Tree.thy Thu May 26 17:51:22 2016 +0200
@@ -5,7 +5,7 @@
Metis example featuring binary trees.
*)
-section {* Metis Example Featuring Binary Trees *}
+section \<open>Metis Example Featuring Binary Trees\<close>
theory Binary_Tree
imports Main
@@ -53,7 +53,7 @@
"append Lf t = t"
| "append (Br a t1 t2) t = Br a (append t1 t) (append t2 t)"
-text {* \medskip BT simplification *}
+text \<open>\medskip BT simplification\<close>
lemma n_leaves_reflect: "n_leaves (reflect t) = n_leaves t"
proof (induct t)
@@ -82,9 +82,9 @@
apply (metis depth.simps(1) reflect.simps(1))
by (metis depth.simps(2) max.commute reflect.simps(2))
-text {*
+text \<open>
The famous relationship between the numbers of leaves and nodes.
-*}
+\<close>
lemma n_leaves_nodes: "n_leaves t = Suc (n_nodes t)"
apply (induct t)
@@ -197,9 +197,9 @@
reflect.simps(1))
by (metis preorder_reflect reflect_reflect_ident rev_swap)
-text {*
+text \<open>
Analogues of the standard properties of the append function for lists.
-*}
+\<close>
lemma append_assoc [simp]: "append (append t1 t2) t3 = append t1 (append t2 t3)"
apply (induct t1)
--- a/src/HOL/Metis_Examples/Clausification.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Metis_Examples/Clausification.thy Thu May 26 17:51:22 2016 +0200
@@ -4,14 +4,14 @@
Example that exercises Metis's Clausifier.
*)
-section {* Example that Exercises Metis's Clausifier *}
+section \<open>Example that Exercises Metis's Clausifier\<close>
theory Clausification
imports Complex_Main
begin
-text {* Definitional CNF for facts *}
+text \<open>Definitional CNF for facts\<close>
declare [[meson_max_clauses = 10]]
@@ -83,7 +83,7 @@
by (metis (full_types) rax)
-text {* Definitional CNF for goal *}
+text \<open>Definitional CNF for goal\<close>
axiomatization p :: "nat \<Rightarrow> nat \<Rightarrow> bool" where
pax: "\<exists>b. \<forall>a. (p b a \<and> p 0 0 \<and> p 1 a) \<or> (p 0 1 \<and> p 1 0 \<and> p a b)"
@@ -109,7 +109,7 @@
by (metis (full_types) pax)
-text {* New Skolemizer *}
+text \<open>New Skolemizer\<close>
declare [[metis_new_skolem]]
--- a/src/HOL/Metis_Examples/Message.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Metis_Examples/Message.thy Thu May 26 17:51:22 2016 +0200
@@ -5,7 +5,7 @@
Metis example featuring message authentication.
*)
-section {* Metis Example Featuring Message Authentication *}
+section \<open>Metis Example Featuring Message Authentication\<close>
theory Message
imports Main
@@ -19,8 +19,8 @@
type_synonym key = nat
consts
- all_symmetric :: bool --{*true if all keys are symmetric*}
- invKey :: "key=>key" --{*inverse of a symmetric key*}
+ all_symmetric :: bool \<comment>\<open>true if all keys are symmetric\<close>
+ invKey :: "key=>key" \<comment>\<open>inverse of a symmetric key\<close>
specification (invKey)
invKey [simp]: "invKey (invKey K) = K"
@@ -28,26 +28,26 @@
by (metis id_apply)
-text{*The inverse of a symmetric key is itself; that of a public key
- is the private key and vice versa*}
+text\<open>The inverse of a symmetric key is itself; that of a public key
+ is the private key and vice versa\<close>
definition symKeys :: "key set" where
"symKeys == {K. invKey K = K}"
-datatype --{*We allow any number of friendly agents*}
+datatype \<comment>\<open>We allow any number of friendly agents\<close>
agent = Server | Friend nat | Spy
datatype
- msg = Agent agent --{*Agent names*}
- | Number nat --{*Ordinary integers, timestamps, ...*}
- | Nonce nat --{*Unguessable nonces*}
- | Key key --{*Crypto keys*}
- | Hash msg --{*Hashing*}
- | MPair msg msg --{*Compound messages*}
- | Crypt key msg --{*Encryption, public- or shared-key*}
+ msg = Agent agent \<comment>\<open>Agent names\<close>
+ | Number nat \<comment>\<open>Ordinary integers, timestamps, ...\<close>
+ | Nonce nat \<comment>\<open>Unguessable nonces\<close>
+ | Key key \<comment>\<open>Crypto keys\<close>
+ | Hash msg \<comment>\<open>Hashing\<close>
+ | MPair msg msg \<comment>\<open>Compound messages\<close>
+ | Crypt key msg \<comment>\<open>Encryption, public- or shared-key\<close>
-text{*Concrete syntax: messages appear as \<open>\<lbrace>A,B,NA\<rbrace>\<close>, etc...*}
+text\<open>Concrete syntax: messages appear as \<open>\<lbrace>A,B,NA\<rbrace>\<close>, etc...\<close>
syntax
"_MTuple" :: "['a, args] => 'a * 'b" ("(2\<lbrace>_,/ _\<rbrace>)")
translations
@@ -56,15 +56,15 @@
definition HPair :: "[msg,msg] => msg" ("(4Hash[_] /_)" [0, 1000]) where
- --{*Message Y paired with a MAC computed with the help of X*}
+ \<comment>\<open>Message Y paired with a MAC computed with the help of X\<close>
"Hash[X] Y == \<lbrace> Hash\<lbrace>X,Y\<rbrace>, Y\<rbrace>"
definition keysFor :: "msg set => key set" where
- --{*Keys useful to decrypt elements of a message set*}
+ \<comment>\<open>Keys useful to decrypt elements of a message set\<close>
"keysFor H == invKey ` {K. \<exists>X. Crypt K X \<in> H}"
-subsubsection{*Inductive Definition of All Parts" of a Message*}
+subsubsection\<open>Inductive Definition of All Parts" of a Message\<close>
inductive_set
parts :: "msg set => msg set"
@@ -83,7 +83,7 @@
apply (metis parts.Snd)
by (metis parts.Body)
-text{*Equations hold because constructors are injective.*}
+text\<open>Equations hold because constructors are injective.\<close>
lemma Friend_image_eq [simp]: "(Friend x \<in> Friend`A) = (x:A)"
by (metis agent.inject image_iff)
@@ -94,13 +94,13 @@
by (metis image_iff msg.distinct(23))
-subsubsection{*Inverse of keys *}
+subsubsection\<open>Inverse of keys\<close>
lemma invKey_eq [simp]: "(invKey K = invKey K') = (K = K')"
by (metis invKey)
-subsection{*keysFor operator*}
+subsection\<open>keysFor operator\<close>
lemma keysFor_empty [simp]: "keysFor {} = {}"
by (unfold keysFor_def, blast)
@@ -111,7 +111,7 @@
lemma keysFor_UN [simp]: "keysFor (\<Union>i\<in>A. H i) = (\<Union>i\<in>A. keysFor (H i))"
by (unfold keysFor_def, blast)
-text{*Monotonicity*}
+text\<open>Monotonicity\<close>
lemma keysFor_mono: "G \<subseteq> H ==> keysFor(G) \<subseteq> keysFor(H)"
by (unfold keysFor_def, blast)
@@ -144,7 +144,7 @@
by (unfold keysFor_def, blast)
-subsection{*Inductive relation "parts"*}
+subsection\<open>Inductive relation "parts"\<close>
lemma MPair_parts:
"[| \<lbrace>X,Y\<rbrace> \<in> parts H;
@@ -152,10 +152,10 @@
by (blast dest: parts.Fst parts.Snd)
declare MPair_parts [elim!] parts.Body [dest!]
-text{*NB These two rules are UNSAFE in the formal sense, as they discard the
+text\<open>NB These two rules are UNSAFE in the formal sense, as they discard the
compound message. They work well on THIS FILE.
- @{text MPair_parts} is left as SAFE because it speeds up proofs.
- The Crypt rule is normally kept UNSAFE to avoid breaking up certificates.*}
+ \<open>MPair_parts\<close> is left as SAFE because it speeds up proofs.
+ The Crypt rule is normally kept UNSAFE to avoid breaking up certificates.\<close>
lemma parts_increasing: "H \<subseteq> parts(H)"
by blast
@@ -171,14 +171,14 @@
lemma parts_emptyE [elim!]: "X\<in> parts{} ==> P"
by simp
-text{*WARNING: loops if H = {Y}, therefore must not be repeated!*}
+text\<open>WARNING: loops if H = {Y}, therefore must not be repeated!\<close>
lemma parts_singleton: "X\<in> parts H ==> \<exists>Y\<in>H. X\<in> parts {Y}"
apply (erule parts.induct)
apply fast+
done
-subsubsection{*Unions *}
+subsubsection\<open>Unions\<close>
lemma parts_Un_subset1: "parts(G) \<union> parts(H) \<subseteq> parts(G \<union> H)"
by (intro Un_least parts_mono Un_upper1 Un_upper2)
@@ -212,19 +212,19 @@
lemma parts_UN [simp]: "parts(\<Union>x\<in>A. H x) = (\<Union>x\<in>A. parts(H x))"
by (intro equalityI parts_UN_subset1 parts_UN_subset2)
-text{*Added to simplify arguments to parts, analz and synth.
- NOTE: the UN versions are no longer used!*}
+text\<open>Added to simplify arguments to parts, analz and synth.
+ NOTE: the UN versions are no longer used!\<close>
-text{*This allows @{text blast} to simplify occurrences of
- @{term "parts(G\<union>H)"} in the assumption.*}
+text\<open>This allows \<open>blast\<close> to simplify occurrences of
+ @{term "parts(G\<union>H)"} in the assumption.\<close>
lemmas in_parts_UnE = parts_Un [THEN equalityD1, THEN subsetD, THEN UnE]
declare in_parts_UnE [elim!]
lemma parts_insert_subset: "insert X (parts H) \<subseteq> parts(insert X H)"
by (blast intro: parts_mono [THEN [2] rev_subsetD])
-subsubsection{*Idempotence and transitivity *}
+subsubsection\<open>Idempotence and transitivity\<close>
lemma parts_partsD [dest!]: "X\<in> parts (parts H) ==> X\<in> parts H"
by (erule parts.induct, blast+)
@@ -245,7 +245,7 @@
by (metis (no_types) Un_insert_left Un_insert_right insert_absorb le_supE
parts_Un parts_idem parts_increasing parts_trans)
-subsubsection{*Rewrite rules for pulling out atomic messages *}
+subsubsection\<open>Rewrite rules for pulling out atomic messages\<close>
lemmas parts_insert_eq_I = equalityI [OF subsetI parts_insert_subset]
@@ -310,11 +310,11 @@
apply (metis le_trans linorder_linear)
done
-subsection{*Inductive relation "analz"*}
+subsection\<open>Inductive relation "analz"\<close>
-text{*Inductive definition of "analz" -- what can be broken down from a set of
+text\<open>Inductive definition of "analz" -- what can be broken down from a set of
messages, including keys. A form of downward closure. Pairs can
- be taken apart; messages decrypted with known keys. *}
+ be taken apart; messages decrypted with known keys.\<close>
inductive_set
analz :: "msg set => msg set"
@@ -327,14 +327,14 @@
"[|Crypt K X \<in> analz H; Key(invKey K): analz H|] ==> X \<in> analz H"
-text{*Monotonicity; Lemma 1 of Lowe's paper*}
+text\<open>Monotonicity; Lemma 1 of Lowe's paper\<close>
lemma analz_mono: "G\<subseteq>H ==> analz(G) \<subseteq> analz(H)"
apply auto
apply (erule analz.induct)
apply (auto dest: analz.Fst analz.Snd)
done
-text{*Making it safe speeds up proofs*}
+text\<open>Making it safe speeds up proofs\<close>
lemma MPair_analz [elim!]:
"[| \<lbrace>X,Y\<rbrace> \<in> analz H;
[| X \<in> analz H; Y \<in> analz H |] ==> P
@@ -367,22 +367,22 @@
lemmas analz_insertI = subset_insertI [THEN analz_mono, THEN [2] rev_subsetD]
-subsubsection{*General equational properties *}
+subsubsection\<open>General equational properties\<close>
lemma analz_empty [simp]: "analz{} = {}"
apply safe
apply (erule analz.induct, blast+)
done
-text{*Converse fails: we can analz more from the union than from the
- separate parts, as a key in one might decrypt a message in the other*}
+text\<open>Converse fails: we can analz more from the union than from the
+ separate parts, as a key in one might decrypt a message in the other\<close>
lemma analz_Un: "analz(G) \<union> analz(H) \<subseteq> analz(G \<union> H)"
by (intro Un_least analz_mono Un_upper1 Un_upper2)
lemma analz_insert: "insert X (analz H) \<subseteq> analz(insert X H)"
by (blast intro: analz_mono [THEN [2] rev_subsetD])
-subsubsection{*Rewrite rules for pulling out atomic messages *}
+subsubsection\<open>Rewrite rules for pulling out atomic messages\<close>
lemmas analz_insert_eq_I = equalityI [OF subsetI analz_insert]
@@ -410,7 +410,7 @@
apply (erule analz.induct, auto)
done
-text{*Can only pull out Keys if they are not needed to decrypt the rest*}
+text\<open>Can only pull out Keys if they are not needed to decrypt the rest\<close>
lemma analz_insert_Key [simp]:
"K \<notin> keysFor (analz H) ==>
analz (insert (Key K) H) = insert (Key K) (analz H)"
@@ -429,7 +429,7 @@
apply (blast intro: analz.Fst analz.Snd)+
done
-text{*Can pull out enCrypted message if the Key is not known*}
+text\<open>Can pull out enCrypted message if the Key is not known\<close>
lemma analz_insert_Crypt:
"Key (invKey K) \<notin> analz H
==> analz (insert (Crypt K X) H) = insert (Crypt K X) (analz H)"
@@ -459,10 +459,10 @@
insert (Crypt K X) (analz (insert X H))"
by (intro equalityI lemma1 lemma2)
-text{*Case analysis: either the message is secure, or it is not! Effective,
-but can cause subgoals to blow up! Use with @{text "if_split"}; apparently
-@{text "split_tac"} does not cope with patterns such as @{term"analz (insert
-(Crypt K X) H)"} *}
+text\<open>Case analysis: either the message is secure, or it is not! Effective,
+but can cause subgoals to blow up! Use with \<open>if_split\<close>; apparently
+\<open>split_tac\<close> does not cope with patterns such as @{term"analz (insert
+(Crypt K X) H)"}\<close>
lemma analz_Crypt_if [simp]:
"analz (insert (Crypt K X) H) =
(if (Key (invKey K) \<in> analz H)
@@ -471,7 +471,7 @@
by (simp add: analz_insert_Crypt analz_insert_Decrypt)
-text{*This rule supposes "for the sake of argument" that we have the key.*}
+text\<open>This rule supposes "for the sake of argument" that we have the key.\<close>
lemma analz_insert_Crypt_subset:
"analz (insert (Crypt K X) H) \<subseteq>
insert (Crypt K X) (analz (insert X H))"
@@ -486,7 +486,7 @@
done
-subsubsection{*Idempotence and transitivity *}
+subsubsection\<open>Idempotence and transitivity\<close>
lemma analz_analzD [dest!]: "X\<in> analz (analz H) ==> X\<in> analz H"
by (erule analz.induct, blast+)
@@ -509,14 +509,14 @@
lemma analz_cut: "[| Y\<in> analz (insert X H); X\<in> analz H |] ==> Y\<in> analz H"
by (metis analz_idem analz_increasing analz_mono insert_absorb insert_mono insert_subset)
-text{*This rewrite rule helps in the simplification of messages that involve
+text\<open>This rewrite rule helps in the simplification of messages that involve
the forwarding of unknown components (X). Without it, removing occurrences
- of X can be very complicated. *}
+ of X can be very complicated.\<close>
lemma analz_insert_eq: "X\<in> analz H ==> analz (insert X H) = analz H"
by (blast intro: analz_cut analz_insertI)
-text{*A congruence rule for "analz" *}
+text\<open>A congruence rule for "analz"\<close>
lemma analz_subset_cong:
"[| analz G \<subseteq> analz G'; analz H \<subseteq> analz H' |]
@@ -535,14 +535,14 @@
"analz H = analz H' ==> analz(insert X H) = analz(insert X H')"
by (force simp only: insert_def intro!: analz_cong)
-text{*If there are no pairs or encryptions then analz does nothing*}
+text\<open>If there are no pairs or encryptions then analz does nothing\<close>
lemma analz_trivial:
"[| \<forall>X Y. \<lbrace>X,Y\<rbrace> \<notin> H; \<forall>X K. Crypt K X \<notin> H |] ==> analz H = H"
apply safe
apply (erule analz.induct, blast+)
done
-text{*These two are obsolete (with a single Spy) but cost little to prove...*}
+text\<open>These two are obsolete (with a single Spy) but cost little to prove...\<close>
lemma analz_UN_analz_lemma:
"X\<in> analz (\<Union>i\<in>A. analz (H i)) ==> X\<in> analz (\<Union>i\<in>A. H i)"
apply (erule analz.induct)
@@ -553,12 +553,12 @@
by (blast intro: analz_UN_analz_lemma analz_mono [THEN [2] rev_subsetD])
-subsection{*Inductive relation "synth"*}
+subsection\<open>Inductive relation "synth"\<close>
-text{*Inductive definition of "synth" -- what can be built up from a set of
+text\<open>Inductive definition of "synth" -- what can be built up from a set of
messages. A form of upward closure. Pairs can be built, messages
encrypted with known keys. Agent names are public domain.
- Numbers can be guessed, but Nonces cannot be. *}
+ Numbers can be guessed, but Nonces cannot be.\<close>
inductive_set
synth :: "msg set => msg set"
@@ -571,12 +571,12 @@
| MPair [intro]: "[|X \<in> synth H; Y \<in> synth H|] ==> \<lbrace>X,Y\<rbrace> \<in> synth H"
| Crypt [intro]: "[|X \<in> synth H; Key(K) \<in> H|] ==> Crypt K X \<in> synth H"
-text{*Monotonicity*}
+text\<open>Monotonicity\<close>
lemma synth_mono: "G\<subseteq>H ==> synth(G) \<subseteq> synth(H)"
by (auto, erule synth.induct, auto)
-text{*NO @{text Agent_synth}, as any Agent name can be synthesized.
- The same holds for @{term Number}*}
+text\<open>NO \<open>Agent_synth\<close>, as any Agent name can be synthesized.
+ The same holds for @{term Number}\<close>
inductive_cases Nonce_synth [elim!]: "Nonce n \<in> synth H"
inductive_cases Key_synth [elim!]: "Key K \<in> synth H"
inductive_cases Hash_synth [elim!]: "Hash X \<in> synth H"
@@ -587,17 +587,17 @@
lemma synth_increasing: "H \<subseteq> synth(H)"
by blast
-subsubsection{*Unions *}
+subsubsection\<open>Unions\<close>
-text{*Converse fails: we can synth more from the union than from the
- separate parts, building a compound message using elements of each.*}
+text\<open>Converse fails: we can synth more from the union than from the
+ separate parts, building a compound message using elements of each.\<close>
lemma synth_Un: "synth(G) \<union> synth(H) \<subseteq> synth(G \<union> H)"
by (intro Un_least synth_mono Un_upper1 Un_upper2)
lemma synth_insert: "insert X (synth H) \<subseteq> synth(insert X H)"
by (metis insert_iff insert_subset subset_insertI synth.Inj synth_mono)
-subsubsection{*Idempotence and transitivity *}
+subsubsection\<open>Idempotence and transitivity\<close>
lemma synth_synthD [dest!]: "X\<in> synth (synth H) ==> X\<in> synth H"
by (erule synth.induct, blast+)
@@ -639,7 +639,7 @@
by (unfold keysFor_def, blast)
-subsubsection{*Combinations of parts, analz and synth *}
+subsubsection\<open>Combinations of parts, analz and synth\<close>
lemma parts_synth [simp]: "parts (synth H) = parts H \<union> synth H"
apply (rule equalityI)
@@ -681,7 +681,7 @@
qed
-subsubsection{*For reasoning about the Fake rule in traces *}
+subsubsection\<open>For reasoning about the Fake rule in traces\<close>
lemma parts_insert_subset_Un: "X \<in> G ==> parts(insert X H) \<subseteq> parts G \<union> parts H"
proof -
--- a/src/HOL/Metis_Examples/Proxies.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Metis_Examples/Proxies.thy Thu May 26 17:51:22 2016 +0200
@@ -5,10 +5,10 @@
rudimentary higher-order reasoning.
*)
-section {*
+section \<open>
Example that Exercises Metis's and Sledgehammer's Logical Symbol Proxies for
Rudimentary Higher-Order Reasoning.
-*}
+\<close>
theory Proxies
imports Type_Encodings
@@ -17,7 +17,7 @@
sledgehammer_params [prover = spass, fact_filter = mepo, timeout = 30, preplay_timeout = 0,
dont_minimize]
-text {* Extensionality and set constants *}
+text \<open>Extensionality and set constants\<close>
lemma plus_1_not_0: "n + (1::nat) \<noteq> 0"
by simp
@@ -57,7 +57,7 @@
sledgehammer [expect = some]
by (metis_exhaust B_def C_def int_le_0_imp_le_1 predicate1I)
-text {* Proxies for logical constants *}
+text \<open>Proxies for logical constants\<close>
lemma "id (op =) x x"
sledgehammer [type_enc = erased, expect = none] (id_apply)
--- a/src/HOL/Metis_Examples/Sets.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Metis_Examples/Sets.thy Thu May 26 17:51:22 2016 +0200
@@ -5,7 +5,7 @@
Metis example featuring typed set theory.
*)
-section {* Metis Example Featuring Typed Set Theory *}
+section \<open>Metis Example Featuring Typed Set Theory\<close>
theory Sets
imports Main
@@ -177,10 +177,10 @@
thus "\<exists>z. S \<subseteq> {z}" by metis
qed
-text {*
+text \<open>
From W. W. Bledsoe and Guohui Feng, SET-VAR. JAR 11 (3), 1993, pages
293-314.
-*}
+\<close>
(* Notes: (1) The numbering doesn't completely agree with the paper.
(2) We must rename set variables to avoid type clashes. *)
--- a/src/HOL/Metis_Examples/Tarski.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Metis_Examples/Tarski.thy Thu May 26 17:51:22 2016 +0200
@@ -5,7 +5,7 @@
Metis example featuring the full theorem of Tarski.
*)
-section {* Metis Example Featuring the Full Theorem of Tarski *}
+section \<open>Metis Example Featuring the Full Theorem of Tarski\<close>
theory Tarski
imports Main "~~/src/HOL/Library/FuncSet"
@@ -117,7 +117,7 @@
x: intY1}
(| pset=intY1, order=induced intY1 r|)"
-subsection {* Partial Order *}
+subsection \<open>Partial Order\<close>
lemma (in PO) PO_imp_refl_on: "refl_on A r"
apply (insert cl_po)
@@ -158,13 +158,13 @@
"S \<subseteq> A ==> (| pset = S, order = induced S r |) \<in> PartialOrder"
apply (simp (no_asm) add: PartialOrder_def)
apply auto
--- {* refl *}
+\<comment> \<open>refl\<close>
apply (simp add: refl_on_def induced_def)
apply (blast intro: reflE)
--- {* antisym *}
+\<comment> \<open>antisym\<close>
apply (simp add: antisym_def induced_def)
apply (blast intro: antisymE)
--- {* trans *}
+\<comment> \<open>trans\<close>
apply (simp add: trans_def induced_def)
apply (blast intro: transE)
done
@@ -290,7 +290,7 @@
apply (simp add: reflE)
done
-subsection {* sublattice *}
+subsection \<open>sublattice\<close>
lemma (in PO) sublattice_imp_CL:
"S <<= cl ==> (| pset = S, order = induced S r |) \<in> CompleteLattice"
@@ -301,7 +301,7 @@
==> S <<= cl"
by (simp add: sublattice_def A_def r_def)
-subsection {* lub *}
+subsection \<open>lub\<close>
lemma (in CL) lub_unique: "[| S \<subseteq> A; isLub S cl x; isLub S cl L|] ==> x = L"
apply (rule antisymE)
@@ -366,7 +366,7 @@
(\<forall>z \<in> A. (\<forall>y \<in> S. (y, z):r) --> (L, z) \<in> r)|] ==> isLub S cl L"
by (simp add: isLub_def A_def r_def)
-subsection {* glb *}
+subsection \<open>glb\<close>
lemma (in CL) glb_in_lattice: "S \<subseteq> A ==> glb S cl \<in> A"
apply (subst glb_dual_lub)
@@ -394,10 +394,10 @@
apply (simp add: dualA_iff A_def, assumption)
done
-text {*
+text \<open>
Reduce the sublattice property by using substructural properties;
- abandoned see @{text "Tarski_4.ML"}.
-*}
+ abandoned see \<open>Tarski_4.ML\<close>.
+\<close>
declare (in CLF) f_cl [simp]
@@ -433,7 +433,7 @@
declare (in CLF) CLF_set_def[simp del] CL_dualCL[simp del] monotone_dual[simp del]
dualA_iff[simp del]
-subsection {* fixed points *}
+subsection \<open>fixed points\<close>
lemma fix_subset: "fix f A \<subseteq> A"
by (simp add: fix_def, fast)
@@ -445,7 +445,7 @@
"[| A \<subseteq> B; x \<in> fix (%y: A. f y) A |] ==> x \<in> fix f B"
by (simp add: fix_def, auto)
-subsection {* lemmas for Tarski, lub *}
+subsection \<open>lemmas for Tarski, lub\<close>
(*never proved, 2007-01-22*)
@@ -456,14 +456,14 @@
apply (rule lub_least, fast)
apply (rule f_in_funcset [THEN funcset_mem])
apply (rule lub_in_lattice, fast)
--- {* @{text "\<forall>x:H. (x, f (lub H r)) \<in> r"} *}
+\<comment> \<open>\<open>\<forall>x:H. (x, f (lub H r)) \<in> r\<close>\<close>
apply (rule ballI)
(*never proved, 2007-01-22*)
apply (rule transE)
--- {* instantiates @{text "(x, ?z) \<in> order cl to (x, f x)"}, *}
--- {* because of the definition of @{text H} *}
+\<comment> \<open>instantiates \<open>(x, ?z) \<in> order cl to (x, f x)\<close>,\<close>
+\<comment> \<open>because of the definition of \<open>H\<close>\<close>
apply fast
--- {* so it remains to show @{text "(f x, f (lub H cl)) \<in> r"} *}
+\<comment> \<open>so it remains to show \<open>(f x, f (lub H cl)) \<in> r\<close>\<close>
apply (rule_tac f = "f" in monotoneE)
apply (rule monotone_f, fast)
apply (rule lub_in_lattice, fast)
@@ -557,7 +557,7 @@
apply (metis P_def lubH_is_fixp)
done
-subsection {* Tarski fixpoint theorem 1, first part *}
+subsection \<open>Tarski fixpoint theorem 1, first part\<close>
declare CL.lubI[intro] fix_subset[intro] CL.lub_in_lattice[intro]
CLF.fixf_le_lubH[simp] CLF.lubH_least_fixf[simp]
@@ -582,7 +582,7 @@
PO.dualPO[intro] CL.CL_dualCL[intro] PO.dualr_iff[simp]
lemma (in CLF) glbH_is_fixp: "H = {x. (f x, x) \<in> r & x \<in> A} ==> glb H cl \<in> P"
- -- {* Tarski for glb *}
+ \<comment> \<open>Tarski for glb\<close>
(*sledgehammer;*)
apply (simp add: glb_dual_lub P_def A_def r_def)
apply (rule dualA_iff [THEN subst])
@@ -613,7 +613,7 @@
OF dualPO CL_dualCL] dualPO CL_dualCL CLF_dual dualr_iff)
done
-subsection {* interval *}
+subsection \<open>interval\<close>
declare (in CLF) CO_refl_on[simp] refl_on_def [simp]
@@ -671,7 +671,7 @@
apply (rule ballI)
apply (simp add: interval_lemma1)
apply (simp add: isLub_upper)
--- {* @{text "(L, b) \<in> r"} *}
+\<comment> \<open>\<open>(L, b) \<in> r\<close>\<close>
apply (simp add: isLub_least interval_lemma2)
done
@@ -709,31 +709,31 @@
prefer 2 apply assumption
apply assumption
apply (erule exE)
--- {* define the lub for the interval as *}
+\<comment> \<open>define the lub for the interval as\<close>
apply (rule_tac x = "if S = {} then a else L" in exI)
apply (simp (no_asm_simp) add: isLub_def split del: if_split)
apply (intro impI conjI)
--- {* @{text "(if S = {} then a else L) \<in> interval r a b"} *}
+\<comment> \<open>\<open>(if S = {} then a else L) \<in> interval r a b\<close>\<close>
apply (simp add: CL_imp_PO L_in_interval)
apply (simp add: left_in_interval)
--- {* lub prop 1 *}
+\<comment> \<open>lub prop 1\<close>
apply (case_tac "S = {}")
--- {* @{text "S = {}, y \<in> S = False => everything"} *}
+\<comment> \<open>\<open>S = {}, y \<in> S = False => everything\<close>\<close>
apply fast
--- {* @{text "S \<noteq> {}"} *}
+\<comment> \<open>\<open>S \<noteq> {}\<close>\<close>
apply simp
--- {* @{text "\<forall>y:S. (y, L) \<in> induced (interval r a b) r"} *}
+\<comment> \<open>\<open>\<forall>y:S. (y, L) \<in> induced (interval r a b) r\<close>\<close>
apply (rule ballI)
apply (simp add: induced_def L_in_interval)
apply (rule conjI)
apply (rule subsetD)
apply (simp add: S_intv_cl, assumption)
apply (simp add: isLub_upper)
--- {* @{text "\<forall>z:interval r a b. (\<forall>y:S. (y, z) \<in> induced (interval r a b) r \<longrightarrow> (if S = {} then a else L, z) \<in> induced (interval r a b) r"} *}
+\<comment> \<open>\<open>\<forall>z:interval r a b. (\<forall>y:S. (y, z) \<in> induced (interval r a b) r \<longrightarrow> (if S = {} then a else L, z) \<in> induced (interval r a b) r\<close>\<close>
apply (rule ballI)
apply (rule impI)
apply (case_tac "S = {}")
--- {* @{text "S = {}"} *}
+\<comment> \<open>\<open>S = {}\<close>\<close>
apply simp
apply (simp add: induced_def interval_def)
apply (rule conjI)
@@ -741,7 +741,7 @@
apply (rule interval_not_empty)
apply (rule CO_trans)
apply (simp add: interval_def)
--- {* @{text "S \<noteq> {}"} *}
+\<comment> \<open>\<open>S \<noteq> {}\<close>\<close>
apply simp
apply (simp add: induced_def L_in_interval)
apply (rule isLub_least, assumption)
@@ -771,7 +771,7 @@
lemmas (in CLF) interv_is_compl_latt =
interval_is_sublattice [THEN sublattice_imp_CL]
-subsection {* Top and Bottom *}
+subsection \<open>Top and Bottom\<close>
lemma (in CLF) Top_dual_Bot: "Top cl = Bot (dual cl)"
by (simp add: Top_def Bot_def least_def greatest_def dualA_iff dualr_iff)
@@ -826,7 +826,7 @@
apply (metis Bot_prop ex_in_conv intervalI reflE rel_imp_elem)
done
-subsection {* fixed points form a partial order *}
+subsection \<open>fixed points form a partial order\<close>
lemma (in CLF) fixf_po: "(| pset = P, order = induced P r|) \<in> PartialOrder"
by (simp add: P_def fix_subset po_subset_po)
@@ -856,12 +856,12 @@
apply (rule Y_subset_A)
apply (rule f_in_funcset [THEN funcset_mem])
apply (rule lubY_in_A)
--- {* @{text "Y \<subseteq> P ==> f x = x"} *}
+\<comment> \<open>\<open>Y \<subseteq> P ==> f x = x\<close>\<close>
apply (rule ballI)
(*sledgehammer *)
apply (rule_tac t = "x" in fix_imp_eq [THEN subst])
apply (erule Y_ss [simplified P_def, THEN subsetD])
--- {* @{text "reduce (f x, f (lub Y cl)) \<in> r to (x, lub Y cl) \<in> r"} by monotonicity *}
+\<comment> \<open>\<open>reduce (f x, f (lub Y cl)) \<in> r to (x, lub Y cl) \<in> r\<close> by monotonicity\<close>
(*sledgehammer*)
apply (rule_tac f = "f" in monotoneE)
apply (rule monotone_f)
@@ -890,14 +890,14 @@
apply (rule conjI)
apply (rule transE)
apply (rule lubY_le_flubY)
--- {* @{text "(f (lub Y cl), f x) \<in> r"} *}
+\<comment> \<open>\<open>(f (lub Y cl), f x) \<in> r\<close>\<close>
(*sledgehammer [has been proved before now...]*)
apply (rule_tac f=f in monotoneE)
apply (rule monotone_f)
apply (rule lubY_in_A)
apply (simp add: intY1_def interval_def intY1_elem)
apply (simp add: intY1_def interval_def)
--- {* @{text "(f x, Top cl) \<in> r"} *}
+\<comment> \<open>\<open>(f x, Top cl) \<in> r\<close>\<close>
apply (rule Top_prop)
apply (rule f_in_funcset [THEN funcset_mem])
apply (simp add: intY1_def interval_def intY1_elem)
@@ -966,12 +966,12 @@
"\<exists>L. isLub Y (| pset = P, order = induced P r |) L"
apply (rule_tac x = "v" in exI)
apply (simp add: isLub_def)
--- {* @{text "v \<in> P"} *}
+\<comment> \<open>\<open>v \<in> P\<close>\<close>
apply (simp add: v_in_P)
apply (rule conjI)
(*sledgehammer*)
--- {* @{text v} is lub *}
--- {* @{text "1. \<forall>y:Y. (y, v) \<in> induced P r"} *}
+\<comment> \<open>\<open>v\<close> is lub\<close>
+\<comment> \<open>\<open>1. \<forall>y:Y. (y, v) \<in> induced P r\<close>\<close>
apply (rule ballI)
apply (simp add: induced_def subsetD v_in_P)
apply (rule conjI)
@@ -984,7 +984,7 @@
apply (fold intY1_def)
apply (rule CL.glb_in_lattice [OF CL.intro, OF PO.intro CL_axioms.intro, OF _ intY1_is_cl, simplified])
apply (simp add: CL_imp_PO intY1_is_cl, force)
--- {* @{text v} is LEAST ub *}
+\<comment> \<open>\<open>v\<close> is LEAST ub\<close>
apply clarify
apply (rule indI)
prefer 3 apply assumption
--- a/src/HOL/Metis_Examples/Trans_Closure.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Metis_Examples/Trans_Closure.thy Thu May 26 17:51:22 2016 +0200
@@ -5,7 +5,7 @@
Metis example featuring the transitive closure.
*)
-section {* Metis Example Featuring the Transitive Closure *}
+section \<open>Metis Example Featuring the Transitive Closure\<close>
theory Trans_Closure
imports Main
@@ -16,11 +16,11 @@
type_synonym addr = nat
datatype val
- = Unit -- "dummy result value of void expressions"
- | Null -- "null reference"
- | Bool bool -- "Boolean value"
- | Intg int -- "integer value"
- | Addr addr -- "addresses of objects in the heap"
+ = Unit \<comment> "dummy result value of void expressions"
+ | Null \<comment> "null reference"
+ | Bool bool \<comment> "Boolean value"
+ | Intg int \<comment> "integer value"
+ | Addr addr \<comment> "addresses of objects in the heap"
consts R :: "(addr \<times> addr) set"
--- a/src/HOL/Metis_Examples/Type_Encodings.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Metis_Examples/Type_Encodings.thy Thu May 26 17:51:22 2016 +0200
@@ -4,9 +4,9 @@
Example that exercises Metis's (and hence Sledgehammer's) type encodings.
*)
-section {*
+section \<open>
Example that Exercises Metis's (and Hence Sledgehammer's) Type Encodings
-*}
+\<close>
theory Type_Encodings
imports Main
@@ -14,11 +14,11 @@
declare [[metis_new_skolem]]
-text {* Setup for testing Metis exhaustively *}
+text \<open>Setup for testing Metis exhaustively\<close>
lemma fork: "P \<Longrightarrow> P \<Longrightarrow> P" by assumption
-ML {*
+ML \<open>
val type_encs =
["erased",
"poly_guards",
@@ -60,14 +60,14 @@
THEN COND (has_fewer_prems 2) all_tac no_tac
THEN tac type_encs)
in tac type_encs end
-*}
+\<close>
-method_setup metis_exhaust = {*
+method_setup metis_exhaust = \<open>
Attrib.thms >>
(fn ths => fn ctxt => SIMPLE_METHOD (metis_exhaust_tac ctxt ths))
-*} "exhaustively run Metis with all type encodings"
+\<close> "exhaustively run Metis with all type encodings"
-text {* Miscellaneous tests *}
+text \<open>Miscellaneous tests\<close>
lemma "x = y \<Longrightarrow> y = x"
by metis_exhaust
--- a/src/HOL/Mirabelle/Mirabelle.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Mirabelle/Mirabelle.thy Thu May 26 17:51:22 2016 +0200
@@ -9,13 +9,13 @@
ML_file "Tools/mirabelle.ML"
ML_file "../TPTP/sledgehammer_tactics.ML"
-ML {* Toplevel.add_hook Mirabelle.step_hook *}
+ML \<open>Toplevel.add_hook Mirabelle.step_hook\<close>
-ML {*
+ML \<open>
signature MIRABELLE_ACTION =
sig
val invoke : (string * string) list -> theory -> theory
end
-*}
+\<close>
end
--- a/src/HOL/Mirabelle/Mirabelle_Test.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Mirabelle/Mirabelle_Test.thy Thu May 26 17:51:22 2016 +0200
@@ -2,7 +2,7 @@
Author: Sascha Boehme, TU Munich
*)
-section {* Simple test theory for Mirabelle actions *}
+section \<open>Simple test theory for Mirabelle actions\<close>
theory Mirabelle_Test
imports Main Mirabelle
@@ -16,9 +16,9 @@
ML_file "Tools/mirabelle_sledgehammer_filter.ML"
ML_file "Tools/mirabelle_try0.ML"
-text {*
+text \<open>
Only perform type-checking of the actions,
any reasonable test would be too complicated.
-*}
+\<close>
end
--- a/src/HOL/Mirabelle/ex/Ex.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Mirabelle/ex/Ex.thy Thu May 26 17:51:22 2016 +0200
@@ -1,12 +1,12 @@
theory Ex imports Pure
begin
-ML {*
+ML \<open>
val rc = Isabelle_System.bash
"cd \"$ISABELLE_HOME/src/HOL/Library\"; isabelle mirabelle arith Inner_Product.thy";
if rc <> 0 then error ("Mirabelle example failed: " ^ string_of_int rc)
else ();
-*} -- "some arbitrary small test case"
+\<close> \<comment> "some arbitrary small test case"
end
--- a/src/HOL/Mutabelle/MutabelleExtra.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Mutabelle/MutabelleExtra.thy Thu May 26 17:51:22 2016 +0200
@@ -15,9 +15,9 @@
ML_file "mutabelle_extra.ML"
-section {* configuration *}
+section \<open>configuration\<close>
-ML {* val log_directory = "" *}
+ML \<open>val log_directory = ""\<close>
quickcheck_params [quiet, finite_types = false, report = false, size = 5, iterations = 1000]
@@ -27,25 +27,25 @@
refute_params [maxtime = 10, minsize = 1, maxsize = 5, satsolver = jerusat]
*)
-ML {* val mtds = [
+ML \<open>val mtds = [
MutabelleExtra.quickcheck_mtd (Config.put Quickcheck.finite_types false) "random",
MutabelleExtra.quickcheck_mtd (Config.put Quickcheck.finite_types true) "random",
MutabelleExtra.quickcheck_mtd (Config.put Quickcheck.finite_types false) "small",
MutabelleExtra.quickcheck_mtd (Config.put Quickcheck.finite_types true) "small"
]
-*}
+\<close>
-ML {*
+ML \<open>
fun mutation_testing_of (name, thy) =
(MutabelleExtra.init_random 1.0;
MutabelleExtra.thms_of false thy
|> MutabelleExtra.take_random 200
|> (fn thms => MutabelleExtra.mutate_theorems_and_write_report
@{theory} (1, 50) mtds thms (log_directory ^ "/" ^ name)))
-*}
+\<close>
-text {* Uncomment the following ML code to check the counterexample generation with all theorems of Complex_Main. *}
+text \<open>Uncomment the following ML code to check the counterexample generation with all theorems of Complex_Main.\<close>
(*
ML {*
@@ -64,9 +64,9 @@
*}
*)
-section {* Mutation testing Isabelle theories *}
+section \<open>Mutation testing Isabelle theories\<close>
-subsection {* List theory *}
+subsection \<open>List theory\<close>
(*
ML {*
@@ -74,9 +74,9 @@
*}
*)
-section {* Mutation testing AFP theories *}
+section \<open>Mutation testing AFP theories\<close>
-subsection {* AVL-Trees *}
+subsection \<open>AVL-Trees\<close>
(*
ML {*
@@ -84,7 +84,7 @@
*}
*)
-subsection {* BinaryTree *}
+subsection \<open>BinaryTree\<close>
(*
ML {*
@@ -92,7 +92,7 @@
*}
*)
-subsection {* Huffman *}
+subsection \<open>Huffman\<close>
(*
ML {*
@@ -100,7 +100,7 @@
*}
*)
-subsection {* List_Index *}
+subsection \<open>List_Index\<close>
(*
ML {*
@@ -108,7 +108,7 @@
*}
*)
-subsection {* Matrix *}
+subsection \<open>Matrix\<close>
(*
ML {*
@@ -116,7 +116,7 @@
*}
*)
-subsection {* NBE *}
+subsection \<open>NBE\<close>
(*
ML {*
@@ -124,7 +124,7 @@
*}
*)
-subsection {* Polynomial *}
+subsection \<open>Polynomial\<close>
(*
ML {*
@@ -132,7 +132,7 @@
*}
*)
-subsection {* Presburger Automata *}
+subsection \<open>Presburger Automata\<close>
(*
ML {*
--- a/src/HOL/NanoJava/AxSem.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/NanoJava/AxSem.thy Thu May 26 17:51:22 2016 +0200
@@ -72,13 +72,13 @@
Impl (D,m) {Q} ==>
A \<turnstile> {P} Meth (C,m) {Q}"
- --{* @{text "\<Union>Z"} instead of @{text "\<forall>Z"} in the conclusion and\\
- Z restricted to type state due to limitations of the inductive package *}
+ \<comment>\<open>\<open>\<Union>Z\<close> instead of \<open>\<forall>Z\<close> in the conclusion and\\
+ Z restricted to type state due to limitations of the inductive package\<close>
| Impl: "\<forall>Z::state. A\<union> (\<Union>Z. (\<lambda>Cm. (P Z Cm, Impl Cm, Q Z Cm))`Ms) |\<turnstile>
(\<lambda>Cm. (P Z Cm, body Cm, Q Z Cm))`Ms ==>
A |\<turnstile> (\<lambda>Cm. (P Z Cm, Impl Cm, Q Z Cm))`Ms"
---{* structural rules *}
+\<comment>\<open>structural rules\<close>
| Asm: " a \<in> A ==> A |\<turnstile> {a}"
@@ -86,12 +86,12 @@
| ConjE: "[|A |\<turnstile> C; c \<in> C |] ==> A |\<turnstile> {c}"
- --{* Z restricted to type state due to limitations of the inductive package *}
+ \<comment>\<open>Z restricted to type state due to limitations of the inductive package\<close>
| Conseq:"[| \<forall>Z::state. A \<turnstile> {P' Z} c {Q' Z};
\<forall>s t. (\<forall>Z. P' Z s --> Q' Z t) --> (P s --> Q t) |] ==>
A \<turnstile> {P} c {Q }"
- --{* Z restricted to type state due to limitations of the inductive package *}
+ \<comment>\<open>Z restricted to type state due to limitations of the inductive package\<close>
| eConseq:"[| \<forall>Z::state. A \<turnstile>\<^sub>e {P' Z} e {Q' Z};
\<forall>s v t. (\<forall>Z. P' Z s --> Q' Z v t) --> (P s --> Q v t) |] ==>
A \<turnstile>\<^sub>e {P} e {Q }"
--- a/src/HOL/NanoJava/Decl.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/NanoJava/Decl.thy Thu May 26 17:51:22 2016 +0200
@@ -8,10 +8,10 @@
theory Decl imports Term begin
datatype ty
- = NT --{* null type *}
- | Class cname --{* class type *}
+ = NT \<comment>\<open>null type\<close>
+ | Class cname \<comment>\<open>class type\<close>
-text{* Field declaration *}
+text\<open>Field declaration\<close>
type_synonym fdecl
= "fname \<times> ty"
@@ -21,7 +21,7 @@
lcl ::"(vname \<times> ty) list"
bdy :: stmt
-text{* Method declaration *}
+text\<open>Method declaration\<close>
type_synonym mdecl
= "mname \<times> methd"
@@ -30,7 +30,7 @@
flds ::"fdecl list"
methods ::"mdecl list"
-text{* Class declaration *}
+text\<open>Class declaration\<close>
type_synonym cdecl
= "cname \<times> class"
@@ -45,9 +45,9 @@
(type) "prog " \<leftharpoondown> (type) "cdecl list"
axiomatization
- Prog :: prog --{* program as a global value *}
+ Prog :: prog \<comment>\<open>program as a global value\<close>
and
- Object :: cname --{* name of root class *}
+ Object :: cname \<comment>\<open>name of root class\<close>
definition "class" :: "cname \<rightharpoonup> class" where
--- a/src/HOL/NanoJava/Equivalence.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/NanoJava/Equivalence.thy Thu May 26 17:51:22 2016 +0200
@@ -94,9 +94,9 @@
apply (tactic "split_all_tac @{context} 1", rename_tac P e Q)
apply (rule hoare_ehoare.induct)
(*18*)
-apply (tactic {* ALLGOALS (REPEAT o dresolve_tac @{context} [@{thm all_conjunct2}, @{thm all3_conjunct2}]) *})
-apply (tactic {* ALLGOALS (REPEAT o Rule_Insts.thin_tac @{context} "hoare _ _" []) *})
-apply (tactic {* ALLGOALS (REPEAT o Rule_Insts.thin_tac @{context} "ehoare _ _" []) *})
+apply (tactic \<open>ALLGOALS (REPEAT o dresolve_tac @{context} [@{thm all_conjunct2}, @{thm all3_conjunct2}])\<close>)
+apply (tactic \<open>ALLGOALS (REPEAT o Rule_Insts.thin_tac @{context} "hoare _ _" [])\<close>)
+apply (tactic \<open>ALLGOALS (REPEAT o Rule_Insts.thin_tac @{context} "ehoare _ _" [])\<close>)
apply (simp_all only: cnvalid1_eq cenvalid_def2)
apply fast
apply fast
--- a/src/HOL/NanoJava/Example.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/NanoJava/Example.thy Thu May 26 17:51:22 2016 +0200
@@ -9,7 +9,7 @@
imports Equivalence
begin
-text {*
+text \<open>
\begin{verbatim}
class Nat {
@@ -39,7 +39,7 @@
}
\end{verbatim}
-*}
+\<close>
axiomatization where
This_neq_Par [simp]: "This \<noteq> Par" and
@@ -62,8 +62,8 @@
one :: expr
where "one == {Nat}new Nat..suc(<>)"
-text {* The following properties could be derived from a more complete
- program model, which we leave out for laziness. *}
+text \<open>The following properties could be derived from a more complete
+ program model, which we leave out for laziness.\<close>
axiomatization where Nat_no_subclasses [simp]: "D \<preceq>C Nat = (D=Nat)"
--- a/src/HOL/NanoJava/State.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/NanoJava/State.thy Thu May 26 17:51:22 2016 +0200
@@ -10,12 +10,12 @@
definition body :: "cname \<times> mname => stmt" where
"body \<equiv> \<lambda>(C,m). bdy (the (method C m))"
-text {* Locations, i.e.\ abstract references to objects *}
+text \<open>Locations, i.e.\ abstract references to objects\<close>
typedecl loc
datatype val
- = Null --{* null reference *}
- | Addr loc --{* address, i.e. location of object *}
+ = Null \<comment>\<open>null reference\<close>
+ | Addr loc \<comment>\<open>address, i.e. location of object\<close>
type_synonym fields
= "(fname \<rightharpoonup> val)"
@@ -30,11 +30,11 @@
definition init_vars :: "('a \<rightharpoonup> 'b) => ('a \<rightharpoonup> val)" where
"init_vars m == map_option (\<lambda>T. Null) o m"
-text {* private: *}
+text \<open>private:\<close>
type_synonym heap = "loc \<rightharpoonup> obj"
type_synonym locals = "vname \<rightharpoonup> val"
-text {* private: *}
+text \<open>private:\<close>
record state
= heap :: heap
locals :: locals
@@ -51,15 +51,15 @@
"init_locs C m s \<equiv> s (| locals := locals s ++
init_vars (map_of (lcl (the (method C m)))) |)"
-text {* The first parameter of @{term set_locs} is of type @{typ state}
- rather than @{typ locals} in order to keep @{typ locals} private.*}
+text \<open>The first parameter of @{term set_locs} is of type @{typ state}
+ rather than @{typ locals} in order to keep @{typ locals} private.\<close>
definition set_locs :: "state => state => state" where
"set_locs s s' \<equiv> s' (| locals := locals s |)"
definition get_local :: "state => vname => val" ("_<_>" [99,0] 99) where
"get_local s x \<equiv> the (locals s x)"
---{* local function: *}
+\<comment>\<open>local function:\<close>
definition get_obj :: "state => loc => obj" where
"get_obj s a \<equiv> the (heap s a)"
@@ -69,7 +69,7 @@
definition get_field :: "state => loc => fname => val" where
"get_field s a f \<equiv> the (snd (get_obj s a) f)"
---{* local function: *}
+\<comment>\<open>local function:\<close>
definition hupd :: "loc => obj => state => state" ("hupd'(_\<mapsto>_')" [10,10] 1000) where
"hupd a obj s \<equiv> s (| heap := ((heap s)(a\<mapsto>obj))|)"
--- a/src/HOL/NanoJava/Term.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/NanoJava/Term.thy Thu May 26 17:51:22 2016 +0200
@@ -6,16 +6,16 @@
theory Term imports Main begin
-typedecl cname --{* class name *}
-typedecl mname --{* method name *}
-typedecl fname --{* field name *}
-typedecl vname --{* variable name *}
+typedecl cname \<comment>\<open>class name\<close>
+typedecl mname \<comment>\<open>method name\<close>
+typedecl fname \<comment>\<open>field name\<close>
+typedecl vname \<comment>\<open>variable name\<close>
axiomatization
- This --{* This pointer *}
- Par --{* method parameter *}
- Res :: vname --{* method result *}
- -- {* Inequality axioms are not required for the meta theory. *}
+ This \<comment>\<open>This pointer\<close>
+ Par \<comment>\<open>method parameter\<close>
+ Res :: vname \<comment>\<open>method result\<close>
+ \<comment> \<open>Inequality axioms are not required for the meta theory.\<close>
(*
where
This_neq_Par [simp]: "This \<noteq> Par"
@@ -24,21 +24,21 @@
*)
datatype stmt
- = Skip --{* empty statement *}
+ = Skip \<comment>\<open>empty statement\<close>
| Comp stmt stmt ("_;; _" [91,90 ] 90)
| Cond expr stmt stmt ("If '(_') _ Else _" [ 3,91,91] 91)
| Loop vname stmt ("While '(_') _" [ 3,91 ] 91)
- | LAss vname expr ("_ :== _" [99, 95] 94) --{* local assignment *}
- | FAss expr fname expr ("_.._:==_" [95,99,95] 94) --{* field assignment *}
- | Meth "cname \<times> mname" --{* virtual method *}
- | Impl "cname \<times> mname" --{* method implementation *}
+ | LAss vname expr ("_ :== _" [99, 95] 94) \<comment>\<open>local assignment\<close>
+ | FAss expr fname expr ("_.._:==_" [95,99,95] 94) \<comment>\<open>field assignment\<close>
+ | Meth "cname \<times> mname" \<comment>\<open>virtual method\<close>
+ | Impl "cname \<times> mname" \<comment>\<open>method implementation\<close>
and expr
- = NewC cname ("new _" [ 99] 95) --{* object creation *}
- | Cast cname expr --{* type cast *}
- | LAcc vname --{* local access *}
- | FAcc expr fname ("_.._" [95,99] 95) --{* field access *}
+ = NewC cname ("new _" [ 99] 95) \<comment>\<open>object creation\<close>
+ | Cast cname expr \<comment>\<open>type cast\<close>
+ | LAcc vname \<comment>\<open>local access\<close>
+ | FAcc expr fname ("_.._" [95,99] 95) \<comment>\<open>field access\<close>
| Call cname expr mname expr
- ("{_}_.._'(_')" [99,95,99,95] 95) --{* method call *}
+ ("{_}_.._'(_')" [99,95,99,95] 95) \<comment>\<open>method call\<close>
end
--- a/src/HOL/NanoJava/TypeRel.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/NanoJava/TypeRel.thy Thu May 26 17:51:22 2016 +0200
@@ -8,7 +8,7 @@
imports Decl
begin
-text{* Direct subclass relation *}
+text\<open>Direct subclass relation\<close>
definition subcls1 :: "(cname \<times> cname) set"
where
@@ -24,7 +24,7 @@
subsection "Declarations and properties not used in the meta theory"
-text{* Widening, viz. method invocation conversion *}
+text\<open>Widening, viz. method invocation conversion\<close>
inductive
widen :: "ty => ty => bool" ("_ \<preceq> _" [71,71] 70)
where
@@ -103,7 +103,7 @@
apply (subst cut_apply, auto intro: subcls1I)
done
---{* Methods of a class, with inheritance and hiding *}
+\<comment>\<open>Methods of a class, with inheritance and hiding\<close>
definition "method" :: "cname => (mname \<rightharpoonup> methd)" where
"method C \<equiv> class_rec C methods"
@@ -115,7 +115,7 @@
done
---{* Fields of a class, with inheritance and hiding *}
+\<comment>\<open>Fields of a class, with inheritance and hiding\<close>
definition field :: "cname => (fname \<rightharpoonup> ty)" where
"field C \<equiv> class_rec C flds"
--- a/src/HOL/Nitpick_Examples/Core_Nits.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Nitpick_Examples/Core_Nits.thy Thu May 26 17:51:22 2016 +0200
@@ -5,7 +5,7 @@
Examples featuring Nitpick's functional core.
*)
-section {* Examples Featuring Nitpick's Functional Core *}
+section \<open>Examples Featuring Nitpick's Functional Core\<close>
theory Core_Nits
imports Main
@@ -14,7 +14,7 @@
nitpick_params [verbose, card = 1-6, unary_ints, max_potential = 0,
sat_solver = MiniSat_JNI, max_threads = 1, timeout = 240]
-subsection {* Curry in a Hurry *}
+subsection \<open>Curry in a Hurry\<close>
lemma "(\<lambda>f x y. (curry o case_prod) f x y) = (\<lambda>f x y. (\<lambda>x. x) f x y)"
nitpick [card = 1-12, expect = none]
@@ -36,7 +36,7 @@
nitpick [card = 1-12, expect = none]
by auto
-subsection {* Representations *}
+subsection \<open>Representations\<close>
lemma "\<exists>f. f = (\<lambda>x. x) \<and> f y = y"
nitpick [expect = none]
@@ -156,7 +156,7 @@
nitpick [card = 8, expect = genuine]
oops
-subsection {* Quantifiers *}
+subsection \<open>Quantifiers\<close>
lemma "x = y"
nitpick [card 'a = 1, expect = none]
@@ -264,7 +264,7 @@
nitpick [expect = none]
by auto
-subsection {* Schematic Variables *}
+subsection \<open>Schematic Variables\<close>
schematic_goal "x = ?x"
nitpick [expect = none]
@@ -290,7 +290,7 @@
nitpick [expect = none]
by auto
-subsection {* Known Constants *}
+subsection \<open>Known Constants\<close>
lemma "x \<equiv> Pure.all \<Longrightarrow> False"
nitpick [card = 1, expect = genuine]
@@ -678,7 +678,7 @@
nitpick [expect = none]
oops
-subsection {* The and Eps *}
+subsection \<open>The and Eps\<close>
lemma "x = The"
nitpick [card = 5, expect = genuine]
@@ -919,7 +919,7 @@
nitpick_params [max_potential = 0]
-subsection {* Destructors and Recursors *}
+subsection \<open>Destructors and Recursors\<close>
lemma "(x::'a) = (case True of True \<Rightarrow> x | False \<Rightarrow> x)"
nitpick [card = 2, expect = none]
--- a/src/HOL/Nitpick_Examples/Datatype_Nits.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Nitpick_Examples/Datatype_Nits.thy Thu May 26 17:51:22 2016 +0200
@@ -5,7 +5,7 @@
Examples featuring Nitpick applied to datatypes.
*)
-section {* Examples Featuring Nitpick Applied to Datatypes *}
+section \<open>Examples Featuring Nitpick Applied to Datatypes\<close>
theory Datatype_Nits
imports Main
--- a/src/HOL/Nitpick_Examples/Hotel_Nits.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Nitpick_Examples/Hotel_Nits.thy Thu May 26 17:51:22 2016 +0200
@@ -5,8 +5,8 @@
Nitpick example based on Tobias Nipkow's hotel key card formalization.
*)
-section {* Nitpick Example Based on Tobias Nipkow's Hotel Key Card
- Formalization *}
+section \<open>Nitpick Example Based on Tobias Nipkow's Hotel Key Card
+ Formalization\<close>
theory Hotel_Nits
imports Main
--- a/src/HOL/Nitpick_Examples/Induct_Nits.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Nitpick_Examples/Induct_Nits.thy Thu May 26 17:51:22 2016 +0200
@@ -5,7 +5,7 @@
Examples featuring Nitpick applied to (co)inductive definitions.
*)
-section {* Examples Featuring Nitpick Applied to (Co)inductive Definitions *}
+section \<open>Examples Featuring Nitpick Applied to (Co)inductive Definitions\<close>
theory Induct_Nits
imports Main
--- a/src/HOL/Nitpick_Examples/Integer_Nits.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Nitpick_Examples/Integer_Nits.thy Thu May 26 17:51:22 2016 +0200
@@ -5,7 +5,7 @@
Examples featuring Nitpick applied to natural numbers and integers.
*)
-section {* Examples Featuring Nitpick Applied to Natural Numbers and Integers *}
+section \<open>Examples Featuring Nitpick Applied to Natural Numbers and Integers\<close>
theory Integer_Nits
imports Main
--- a/src/HOL/Nitpick_Examples/Manual_Nits.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Nitpick_Examples/Manual_Nits.thy Thu May 26 17:51:22 2016 +0200
@@ -5,7 +5,7 @@
Examples from the Nitpick manual.
*)
-section {* Examples from the Nitpick Manual *}
+section \<open>Examples from the Nitpick Manual\<close>
(* The "expect" arguments to Nitpick in this theory and the other example
theories are there so that the example can also serve as a regression test
@@ -15,12 +15,12 @@
imports Real "~~/src/HOL/Library/Quotient_Product"
begin
-section {* 2. First Steps *}
+section \<open>2. First Steps\<close>
nitpick_params [sat_solver = MiniSat_JNI, max_threads = 1, timeout = 240]
-subsection {* 2.1. Propositional Logic *}
+subsection \<open>2.1. Propositional Logic\<close>
lemma "P \<longleftrightarrow> Q"
nitpick [expect = genuine]
@@ -30,14 +30,14 @@
oops
-subsection {* 2.2. Type Variables *}
+subsection \<open>2.2. Type Variables\<close>
lemma "x \<in> A \<Longrightarrow> (THE y. y \<in> A) \<in> A"
nitpick [verbose, expect = genuine]
oops
-subsection {* 2.3. Constants *}
+subsection \<open>2.3. Constants\<close>
lemma "x \<in> A \<Longrightarrow> (THE y. y \<in> A) \<in> A"
nitpick [show_consts, expect = genuine]
@@ -51,7 +51,7 @@
by (metis the_equality)
-subsection {* 2.4. Skolemization *}
+subsection \<open>2.4. Skolemization\<close>
lemma "\<exists>g. \<forall>x. g (f x) = x \<Longrightarrow> \<forall>y. \<exists>x. y = f x"
nitpick [expect = genuine]
@@ -66,7 +66,7 @@
oops
-subsection {* 2.5. Natural Numbers and Integers *}
+subsection \<open>2.5. Natural Numbers and Integers\<close>
lemma "\<lbrakk>i \<le> j; n \<le> (m::int)\<rbrakk> \<Longrightarrow> i * n + j * m \<le> i * m + j * n"
nitpick [expect = genuine]
@@ -87,7 +87,7 @@
oops
-subsection {* 2.6. Inductive Datatypes *}
+subsection \<open>2.6. Inductive Datatypes\<close>
lemma "hd (xs @ [y, y]) = hd xs"
nitpick [expect = genuine]
@@ -99,7 +99,7 @@
oops
-subsection {* 2.7. Typedefs, Records, Rationals, and Reals *}
+subsection \<open>2.7. Typedefs, Records, Rationals, and Reals\<close>
definition "three = {0::nat, 1, 2}"
typedef three = three
@@ -129,16 +129,16 @@
nitpick [show_types, expect = genuine]
oops
-ML {*
+ML \<open>
fun my_int_postproc _ _ _ T (Const _ $ (Const _ $ t1 $ t2)) =
HOLogic.mk_number T (snd (HOLogic.dest_number t1)
- snd (HOLogic.dest_number t2))
| my_int_postproc _ _ _ _ t = t
-*}
+\<close>
-declaration {*
+declaration \<open>
Nitpick_Model.register_term_postprocessor @{typ my_int} my_int_postproc
-*}
+\<close>
lemma "add x y = add x x"
nitpick [show_types]
@@ -157,7 +157,7 @@
oops
-subsection {* 2.8. Inductive and Coinductive Predicates *}
+subsection \<open>2.8. Inductive and Coinductive Predicates\<close>
inductive even where
"even 0" |
@@ -200,7 +200,7 @@
oops
-subsection {* 2.9. Coinductive Datatypes *}
+subsection \<open>2.9. Coinductive Datatypes\<close>
codatatype 'a llist = LNil | LCons 'a "'a llist"
@@ -221,7 +221,7 @@
sorry
-subsection {* 2.10. Boxing *}
+subsection \<open>2.10. Boxing\<close>
datatype tm = Var nat | Lam tm | App tm tm
@@ -258,7 +258,7 @@
sorry
-subsection {* 2.11. Scope Monotonicity *}
+subsection \<open>2.11. Scope Monotonicity\<close>
lemma "length xs = length ys \<Longrightarrow> rev (zip xs ys) = zip xs (rev ys)"
nitpick [verbose, expect = genuine]
@@ -270,7 +270,7 @@
oops
-subsection {* 2.12. Inductive Properties *}
+subsection \<open>2.12. Inductive Properties\<close>
inductive_set reach where
"(4::nat) \<in> reach" |
@@ -299,12 +299,12 @@
by (induct set: reach) arith+
-section {* 3. Case Studies *}
+section \<open>3. Case Studies\<close>
nitpick_params [max_potential = 0]
-subsection {* 3.1. A Context-Free Grammar *}
+subsection \<open>3.1. A Context-Free Grammar\<close>
datatype alphabet = a | b
@@ -379,7 +379,7 @@
sorry
-subsection {* 3.2. AA Trees *}
+subsection \<open>3.2. AA Trees\<close>
datatype 'a aa_tree = \<Lambda> | N "'a::linorder" nat "'a aa_tree" "'a aa_tree"
--- a/src/HOL/Nitpick_Examples/Mini_Nits.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Nitpick_Examples/Mini_Nits.thy Thu May 26 17:51:22 2016 +0200
@@ -5,7 +5,7 @@
Examples featuring Minipick, the minimalistic version of Nitpick.
*)
-section {* Examples Featuring Minipick, the Minimalistic Version of Nitpick *}
+section \<open>Examples Featuring Minipick, the Minimalistic Version of Nitpick\<close>
theory Mini_Nits
imports Main
@@ -16,96 +16,96 @@
nitpick_params [verbose, sat_solver = MiniSat_JNI, max_threads = 1,
total_consts = smart]
-ML {*
+ML \<open>
val check = Minipick.minipick @{context}
val expect = Minipick.minipick_expect @{context}
val none = expect "none"
val genuine = expect "genuine"
val unknown = expect "unknown"
-*}
+\<close>
-ML_val {* genuine 1 @{prop "x = Not"} *}
-ML_val {* none 1 @{prop "\<exists>x. x = Not"} *}
-ML_val {* none 1 @{prop "\<not> False"} *}
-ML_val {* genuine 1 @{prop "\<not> True"} *}
-ML_val {* none 1 @{prop "\<not> \<not> b \<longleftrightarrow> b"} *}
-ML_val {* none 1 @{prop True} *}
-ML_val {* genuine 1 @{prop False} *}
-ML_val {* genuine 1 @{prop "True \<longleftrightarrow> False"} *}
-ML_val {* none 1 @{prop "True \<longleftrightarrow> \<not> False"} *}
-ML_val {* none 4 @{prop "\<forall>x. x = x"} *}
-ML_val {* none 4 @{prop "\<exists>x. x = x"} *}
-ML_val {* none 1 @{prop "\<forall>x. x = y"} *}
-ML_val {* genuine 2 @{prop "\<forall>x. x = y"} *}
-ML_val {* none 2 @{prop "\<exists>x. x = y"} *}
-ML_val {* none 2 @{prop "\<forall>x::'a \<times> 'a. x = x"} *}
-ML_val {* none 2 @{prop "\<exists>x::'a \<times> 'a. x = y"} *}
-ML_val {* genuine 2 @{prop "\<forall>x::'a \<times> 'a. x = y"} *}
-ML_val {* none 2 @{prop "\<exists>x::'a \<times> 'a. x = y"} *}
-ML_val {* none 1 @{prop "All = Ex"} *}
-ML_val {* genuine 2 @{prop "All = Ex"} *}
-ML_val {* none 1 @{prop "All P = Ex P"} *}
-ML_val {* genuine 2 @{prop "All P = Ex P"} *}
-ML_val {* none 4 @{prop "x = y \<longrightarrow> P x = P y"} *}
-ML_val {* none 4 @{prop "(x::'a \<times> 'a) = y \<longrightarrow> P x = P y"} *}
-ML_val {* none 2 @{prop "(x::'a \<times> 'a) = y \<longrightarrow> P x y = P y x"} *}
-ML_val {* none 4 @{prop "\<exists>x::'a \<times> 'a. x = y \<longrightarrow> P x = P y"} *}
-ML_val {* none 2 @{prop "(x::'a \<Rightarrow> 'a) = y \<longrightarrow> P x = P y"} *}
-ML_val {* none 2 @{prop "\<exists>x::'a \<Rightarrow> 'a. x = y \<longrightarrow> P x = P y"} *}
-ML_val {* genuine 1 @{prop "(op =) X = Ex"} *}
-ML_val {* none 2 @{prop "\<forall>x::'a \<Rightarrow> 'a. x = x"} *}
-ML_val {* none 1 @{prop "x = y"} *}
-ML_val {* genuine 1 @{prop "x \<longleftrightarrow> y"} *}
-ML_val {* genuine 2 @{prop "x = y"} *}
-ML_val {* genuine 1 @{prop "X \<subseteq> Y"} *}
-ML_val {* none 1 @{prop "P \<and> Q \<longleftrightarrow> Q \<and> P"} *}
-ML_val {* none 1 @{prop "P \<and> Q \<longrightarrow> P"} *}
-ML_val {* none 1 @{prop "P \<or> Q \<longleftrightarrow> Q \<or> P"} *}
-ML_val {* genuine 1 @{prop "P \<or> Q \<longrightarrow> P"} *}
-ML_val {* none 1 @{prop "(P \<longrightarrow> Q) \<longleftrightarrow> (\<not> P \<or> Q)"} *}
-ML_val {* none 4 @{prop "{a} = {a, a}"} *}
-ML_val {* genuine 2 @{prop "{a} = {a, b}"} *}
-ML_val {* genuine 1 @{prop "{a} \<noteq> {a, b}"} *}
-ML_val {* none 4 @{prop "{}\<^sup>+ = {}"} *}
-ML_val {* none 4 @{prop "UNIV\<^sup>+ = UNIV"} *}
-ML_val {* none 4 @{prop "(UNIV :: ('a \<times> 'b) set) - {} = UNIV"} *}
-ML_val {* none 4 @{prop "{} - (UNIV :: ('a \<times> 'b) set) = {}"} *}
-ML_val {* none 1 @{prop "{(a, b), (b, c)}\<^sup>+ = {(a, b), (a, c), (b, c)}"} *}
-ML_val {* genuine 2 @{prop "{(a, b), (b, c)}\<^sup>+ = {(a, b), (a, c), (b, c)}"} *}
-ML_val {* none 4 @{prop "a \<noteq> c \<Longrightarrow> {(a, b), (b, c)}\<^sup>+ = {(a, b), (a, c), (b, c)}"} *}
-ML_val {* none 4 @{prop "A \<union> B = {x. x \<in> A \<or> x \<in> B}"} *}
-ML_val {* none 4 @{prop "A \<inter> B = {x. x \<in> A \<and> x \<in> B}"} *}
-ML_val {* none 4 @{prop "A - B = (\<lambda>x. A x \<and> \<not> B x)"} *}
-ML_val {* none 4 @{prop "\<exists>a b. (a, b) = (b, a)"} *}
-ML_val {* genuine 2 @{prop "(a, b) = (b, a)"} *}
-ML_val {* genuine 2 @{prop "(a, b) \<noteq> (b, a)"} *}
-ML_val {* none 4 @{prop "\<exists>a b::'a \<times> 'a. (a, b) = (b, a)"} *}
-ML_val {* genuine 2 @{prop "(a::'a \<times> 'a, b) = (b, a)"} *}
-ML_val {* none 4 @{prop "\<exists>a b::'a \<times> 'a \<times> 'a. (a, b) = (b, a)"} *}
-ML_val {* genuine 2 @{prop "(a::'a \<times> 'a \<times> 'a, b) \<noteq> (b, a)"} *}
-ML_val {* none 4 @{prop "\<exists>a b::'a \<Rightarrow> 'a. (a, b) = (b, a)"} *}
-ML_val {* genuine 1 @{prop "(a::'a \<Rightarrow> 'a, b) \<noteq> (b, a)"} *}
-ML_val {* none 4 @{prop "fst (a, b) = a"} *}
-ML_val {* none 1 @{prop "fst (a, b) = b"} *}
-ML_val {* genuine 2 @{prop "fst (a, b) = b"} *}
-ML_val {* genuine 2 @{prop "fst (a, b) \<noteq> b"} *}
-ML_val {* genuine 2 @{prop "f ((x, z), y) = (x, z)"} *}
-ML_val {* none 2 @{prop "(ALL x. f x = fst x) \<longrightarrow> f ((x, z), y) = (x, z)"} *}
-ML_val {* none 4 @{prop "snd (a, b) = b"} *}
-ML_val {* none 1 @{prop "snd (a, b) = a"} *}
-ML_val {* genuine 2 @{prop "snd (a, b) = a"} *}
-ML_val {* genuine 2 @{prop "snd (a, b) \<noteq> a"} *}
-ML_val {* genuine 1 @{prop P} *}
-ML_val {* genuine 1 @{prop "(\<lambda>x. P) a"} *}
-ML_val {* genuine 1 @{prop "(\<lambda>x y z. P y x z) a b c"} *}
-ML_val {* none 4 @{prop "\<exists>f. f = (\<lambda>x. x) \<and> f y = y"} *}
-ML_val {* genuine 1 @{prop "\<exists>f. f p \<noteq> p \<and> (\<forall>a b. f (a, b) = (a, b))"} *}
-ML_val {* none 2 @{prop "\<exists>f. \<forall>a b. f (a, b) = (a, b)"} *}
-ML_val {* none 3 @{prop "f = (\<lambda>a b. (b, a)) \<longrightarrow> f x y = (y, x)"} *}
-ML_val {* genuine 2 @{prop "f = (\<lambda>a b. (b, a)) \<longrightarrow> f x y = (x, y)"} *}
-ML_val {* none 4 @{prop "f = (\<lambda>x. f x)"} *}
-ML_val {* none 4 @{prop "f = (\<lambda>x. f x::'a \<Rightarrow> bool)"} *}
-ML_val {* none 4 @{prop "f = (\<lambda>x y. f x y)"} *}
-ML_val {* none 4 @{prop "f = (\<lambda>x y. f x y::'a \<Rightarrow> bool)"} *}
+ML_val \<open>genuine 1 @{prop "x = Not"}\<close>
+ML_val \<open>none 1 @{prop "\<exists>x. x = Not"}\<close>
+ML_val \<open>none 1 @{prop "\<not> False"}\<close>
+ML_val \<open>genuine 1 @{prop "\<not> True"}\<close>
+ML_val \<open>none 1 @{prop "\<not> \<not> b \<longleftrightarrow> b"}\<close>
+ML_val \<open>none 1 @{prop True}\<close>
+ML_val \<open>genuine 1 @{prop False}\<close>
+ML_val \<open>genuine 1 @{prop "True \<longleftrightarrow> False"}\<close>
+ML_val \<open>none 1 @{prop "True \<longleftrightarrow> \<not> False"}\<close>
+ML_val \<open>none 4 @{prop "\<forall>x. x = x"}\<close>
+ML_val \<open>none 4 @{prop "\<exists>x. x = x"}\<close>
+ML_val \<open>none 1 @{prop "\<forall>x. x = y"}\<close>
+ML_val \<open>genuine 2 @{prop "\<forall>x. x = y"}\<close>
+ML_val \<open>none 2 @{prop "\<exists>x. x = y"}\<close>
+ML_val \<open>none 2 @{prop "\<forall>x::'a \<times> 'a. x = x"}\<close>
+ML_val \<open>none 2 @{prop "\<exists>x::'a \<times> 'a. x = y"}\<close>
+ML_val \<open>genuine 2 @{prop "\<forall>x::'a \<times> 'a. x = y"}\<close>
+ML_val \<open>none 2 @{prop "\<exists>x::'a \<times> 'a. x = y"}\<close>
+ML_val \<open>none 1 @{prop "All = Ex"}\<close>
+ML_val \<open>genuine 2 @{prop "All = Ex"}\<close>
+ML_val \<open>none 1 @{prop "All P = Ex P"}\<close>
+ML_val \<open>genuine 2 @{prop "All P = Ex P"}\<close>
+ML_val \<open>none 4 @{prop "x = y \<longrightarrow> P x = P y"}\<close>
+ML_val \<open>none 4 @{prop "(x::'a \<times> 'a) = y \<longrightarrow> P x = P y"}\<close>
+ML_val \<open>none 2 @{prop "(x::'a \<times> 'a) = y \<longrightarrow> P x y = P y x"}\<close>
+ML_val \<open>none 4 @{prop "\<exists>x::'a \<times> 'a. x = y \<longrightarrow> P x = P y"}\<close>
+ML_val \<open>none 2 @{prop "(x::'a \<Rightarrow> 'a) = y \<longrightarrow> P x = P y"}\<close>
+ML_val \<open>none 2 @{prop "\<exists>x::'a \<Rightarrow> 'a. x = y \<longrightarrow> P x = P y"}\<close>
+ML_val \<open>genuine 1 @{prop "(op =) X = Ex"}\<close>
+ML_val \<open>none 2 @{prop "\<forall>x::'a \<Rightarrow> 'a. x = x"}\<close>
+ML_val \<open>none 1 @{prop "x = y"}\<close>
+ML_val \<open>genuine 1 @{prop "x \<longleftrightarrow> y"}\<close>
+ML_val \<open>genuine 2 @{prop "x = y"}\<close>
+ML_val \<open>genuine 1 @{prop "X \<subseteq> Y"}\<close>
+ML_val \<open>none 1 @{prop "P \<and> Q \<longleftrightarrow> Q \<and> P"}\<close>
+ML_val \<open>none 1 @{prop "P \<and> Q \<longrightarrow> P"}\<close>
+ML_val \<open>none 1 @{prop "P \<or> Q \<longleftrightarrow> Q \<or> P"}\<close>
+ML_val \<open>genuine 1 @{prop "P \<or> Q \<longrightarrow> P"}\<close>
+ML_val \<open>none 1 @{prop "(P \<longrightarrow> Q) \<longleftrightarrow> (\<not> P \<or> Q)"}\<close>
+ML_val \<open>none 4 @{prop "{a} = {a, a}"}\<close>
+ML_val \<open>genuine 2 @{prop "{a} = {a, b}"}\<close>
+ML_val \<open>genuine 1 @{prop "{a} \<noteq> {a, b}"}\<close>
+ML_val \<open>none 4 @{prop "{}\<^sup>+ = {}"}\<close>
+ML_val \<open>none 4 @{prop "UNIV\<^sup>+ = UNIV"}\<close>
+ML_val \<open>none 4 @{prop "(UNIV :: ('a \<times> 'b) set) - {} = UNIV"}\<close>
+ML_val \<open>none 4 @{prop "{} - (UNIV :: ('a \<times> 'b) set) = {}"}\<close>
+ML_val \<open>none 1 @{prop "{(a, b), (b, c)}\<^sup>+ = {(a, b), (a, c), (b, c)}"}\<close>
+ML_val \<open>genuine 2 @{prop "{(a, b), (b, c)}\<^sup>+ = {(a, b), (a, c), (b, c)}"}\<close>
+ML_val \<open>none 4 @{prop "a \<noteq> c \<Longrightarrow> {(a, b), (b, c)}\<^sup>+ = {(a, b), (a, c), (b, c)}"}\<close>
+ML_val \<open>none 4 @{prop "A \<union> B = {x. x \<in> A \<or> x \<in> B}"}\<close>
+ML_val \<open>none 4 @{prop "A \<inter> B = {x. x \<in> A \<and> x \<in> B}"}\<close>
+ML_val \<open>none 4 @{prop "A - B = (\<lambda>x. A x \<and> \<not> B x)"}\<close>
+ML_val \<open>none 4 @{prop "\<exists>a b. (a, b) = (b, a)"}\<close>
+ML_val \<open>genuine 2 @{prop "(a, b) = (b, a)"}\<close>
+ML_val \<open>genuine 2 @{prop "(a, b) \<noteq> (b, a)"}\<close>
+ML_val \<open>none 4 @{prop "\<exists>a b::'a \<times> 'a. (a, b) = (b, a)"}\<close>
+ML_val \<open>genuine 2 @{prop "(a::'a \<times> 'a, b) = (b, a)"}\<close>
+ML_val \<open>none 4 @{prop "\<exists>a b::'a \<times> 'a \<times> 'a. (a, b) = (b, a)"}\<close>
+ML_val \<open>genuine 2 @{prop "(a::'a \<times> 'a \<times> 'a, b) \<noteq> (b, a)"}\<close>
+ML_val \<open>none 4 @{prop "\<exists>a b::'a \<Rightarrow> 'a. (a, b) = (b, a)"}\<close>
+ML_val \<open>genuine 1 @{prop "(a::'a \<Rightarrow> 'a, b) \<noteq> (b, a)"}\<close>
+ML_val \<open>none 4 @{prop "fst (a, b) = a"}\<close>
+ML_val \<open>none 1 @{prop "fst (a, b) = b"}\<close>
+ML_val \<open>genuine 2 @{prop "fst (a, b) = b"}\<close>
+ML_val \<open>genuine 2 @{prop "fst (a, b) \<noteq> b"}\<close>
+ML_val \<open>genuine 2 @{prop "f ((x, z), y) = (x, z)"}\<close>
+ML_val \<open>none 2 @{prop "(ALL x. f x = fst x) \<longrightarrow> f ((x, z), y) = (x, z)"}\<close>
+ML_val \<open>none 4 @{prop "snd (a, b) = b"}\<close>
+ML_val \<open>none 1 @{prop "snd (a, b) = a"}\<close>
+ML_val \<open>genuine 2 @{prop "snd (a, b) = a"}\<close>
+ML_val \<open>genuine 2 @{prop "snd (a, b) \<noteq> a"}\<close>
+ML_val \<open>genuine 1 @{prop P}\<close>
+ML_val \<open>genuine 1 @{prop "(\<lambda>x. P) a"}\<close>
+ML_val \<open>genuine 1 @{prop "(\<lambda>x y z. P y x z) a b c"}\<close>
+ML_val \<open>none 4 @{prop "\<exists>f. f = (\<lambda>x. x) \<and> f y = y"}\<close>
+ML_val \<open>genuine 1 @{prop "\<exists>f. f p \<noteq> p \<and> (\<forall>a b. f (a, b) = (a, b))"}\<close>
+ML_val \<open>none 2 @{prop "\<exists>f. \<forall>a b. f (a, b) = (a, b)"}\<close>
+ML_val \<open>none 3 @{prop "f = (\<lambda>a b. (b, a)) \<longrightarrow> f x y = (y, x)"}\<close>
+ML_val \<open>genuine 2 @{prop "f = (\<lambda>a b. (b, a)) \<longrightarrow> f x y = (x, y)"}\<close>
+ML_val \<open>none 4 @{prop "f = (\<lambda>x. f x)"}\<close>
+ML_val \<open>none 4 @{prop "f = (\<lambda>x. f x::'a \<Rightarrow> bool)"}\<close>
+ML_val \<open>none 4 @{prop "f = (\<lambda>x y. f x y)"}\<close>
+ML_val \<open>none 4 @{prop "f = (\<lambda>x y. f x y::'a \<Rightarrow> bool)"}\<close>
end
--- a/src/HOL/Nitpick_Examples/Mono_Nits.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Nitpick_Examples/Mono_Nits.thy Thu May 26 17:51:22 2016 +0200
@@ -5,7 +5,7 @@
Examples featuring Nitpick's monotonicity check.
*)
-section {* Examples Featuring Nitpick's Monotonicity Check *}
+section \<open>Examples Featuring Nitpick's Monotonicity Check\<close>
theory Mono_Nits
imports Main
@@ -13,7 +13,7 @@
(* "~/afp/thys/AVL-Trees/AVL2" "~/afp/thys/Huffman/Huffman" *)
begin
-ML {*
+ML \<open>
open Nitpick_Util
open Nitpick_HOL
open Nitpick_Preproc
@@ -62,84 +62,84 @@
fun nonmono t = not (is_mono t) orelse raise BUG
fun const t = is_const t orelse raise BUG
fun nonconst t = not (is_const t) orelse raise BUG
-*}
+\<close>
-ML {* Nitpick_Mono.trace := false *}
+ML \<open>Nitpick_Mono.trace := false\<close>
-ML_val {* const @{term "A::('a\<Rightarrow>'b)"} *}
-ML_val {* const @{term "(A::'a set) = A"} *}
-ML_val {* const @{term "(A::'a set set) = A"} *}
-ML_val {* const @{term "(\<lambda>x::'a set. a \<in> x)"} *}
-ML_val {* const @{term "{{a::'a}} = C"} *}
-ML_val {* const @{term "{f::'a\<Rightarrow>nat} = {g::'a\<Rightarrow>nat}"} *}
-ML_val {* const @{term "A \<union> (B::'a set)"} *}
-ML_val {* const @{term "\<lambda>A B x::'a. A x \<or> B x"} *}
-ML_val {* const @{term "P (a::'a)"} *}
-ML_val {* const @{term "\<lambda>a::'a. b (c (d::'a)) (e::'a) (f::'a)"} *}
-ML_val {* const @{term "\<forall>A::'a set. a \<in> A"} *}
-ML_val {* const @{term "\<forall>A::'a set. P A"} *}
-ML_val {* const @{term "P \<or> Q"} *}
-ML_val {* const @{term "A \<union> B = (C::'a set)"} *}
-ML_val {* const @{term "(\<lambda>A B x::'a. A x \<or> B x) A B = C"} *}
-ML_val {* const @{term "(if P then (A::'a set) else B) = C"} *}
-ML_val {* const @{term "let A = (C::'a set) in A \<union> B"} *}
-ML_val {* const @{term "THE x::'b. P x"} *}
-ML_val {* const @{term "(\<lambda>x::'a. False)"} *}
-ML_val {* const @{term "(\<lambda>x::'a. True)"} *}
-ML_val {* const @{term "(\<lambda>x::'a. False) = (\<lambda>x::'a. False)"} *}
-ML_val {* const @{term "(\<lambda>x::'a. True) = (\<lambda>x::'a. True)"} *}
-ML_val {* const @{term "Let (a::'a) A"} *}
-ML_val {* const @{term "A (a::'a)"} *}
-ML_val {* const @{term "insert (a::'a) A = B"} *}
-ML_val {* const @{term "- (A::'a set)"} *}
-ML_val {* const @{term "finite (A::'a set)"} *}
-ML_val {* const @{term "\<not> finite (A::'a set)"} *}
-ML_val {* const @{term "finite (A::'a set set)"} *}
-ML_val {* const @{term "\<lambda>a::'a. A a \<and> \<not> B a"} *}
-ML_val {* const @{term "A < (B::'a set)"} *}
-ML_val {* const @{term "A \<le> (B::'a set)"} *}
-ML_val {* const @{term "[a::'a]"} *}
-ML_val {* const @{term "[a::'a set]"} *}
-ML_val {* const @{term "[A \<union> (B::'a set)]"} *}
-ML_val {* const @{term "[A \<union> (B::'a set)] = [C]"} *}
-ML_val {* const @{term "{(\<lambda>x::'a. x = a)} = C"} *}
-ML_val {* const @{term "(\<lambda>a::'a. \<not> A a) = B"} *}
-ML_val {* const @{prop "\<forall>F f g (h::'a set). F f \<and> F g \<and> \<not> f a \<and> g a \<longrightarrow> \<not> f a"} *}
-ML_val {* const @{term "\<lambda>A B x::'a. A x \<and> B x \<and> A = B"} *}
-ML_val {* const @{term "p = (\<lambda>(x::'a) (y::'a). P x \<or> \<not> Q y)"} *}
-ML_val {* const @{term "p = (\<lambda>(x::'a) (y::'a). p x y :: bool)"} *}
-ML_val {* const @{term "p = (\<lambda>A B x. A x \<and> \<not> B x) (\<lambda>x. True) (\<lambda>y. x \<noteq> y)"} *}
-ML_val {* const @{term "p = (\<lambda>y. x \<noteq> y)"} *}
-ML_val {* const @{term "(\<lambda>x. (p::'a\<Rightarrow>bool\<Rightarrow>bool) x False)"} *}
-ML_val {* const @{term "(\<lambda>x y. (p::'a\<Rightarrow>'a\<Rightarrow>bool\<Rightarrow>bool) x y False)"} *}
-ML_val {* const @{term "f = (\<lambda>x::'a. P x \<longrightarrow> Q x)"} *}
-ML_val {* const @{term "\<forall>a::'a. P a"} *}
+ML_val \<open>const @{term "A::('a\<Rightarrow>'b)"}\<close>
+ML_val \<open>const @{term "(A::'a set) = A"}\<close>
+ML_val \<open>const @{term "(A::'a set set) = A"}\<close>
+ML_val \<open>const @{term "(\<lambda>x::'a set. a \<in> x)"}\<close>
+ML_val \<open>const @{term "{{a::'a}} = C"}\<close>
+ML_val \<open>const @{term "{f::'a\<Rightarrow>nat} = {g::'a\<Rightarrow>nat}"}\<close>
+ML_val \<open>const @{term "A \<union> (B::'a set)"}\<close>
+ML_val \<open>const @{term "\<lambda>A B x::'a. A x \<or> B x"}\<close>
+ML_val \<open>const @{term "P (a::'a)"}\<close>
+ML_val \<open>const @{term "\<lambda>a::'a. b (c (d::'a)) (e::'a) (f::'a)"}\<close>
+ML_val \<open>const @{term "\<forall>A::'a set. a \<in> A"}\<close>
+ML_val \<open>const @{term "\<forall>A::'a set. P A"}\<close>
+ML_val \<open>const @{term "P \<or> Q"}\<close>
+ML_val \<open>const @{term "A \<union> B = (C::'a set)"}\<close>
+ML_val \<open>const @{term "(\<lambda>A B x::'a. A x \<or> B x) A B = C"}\<close>
+ML_val \<open>const @{term "(if P then (A::'a set) else B) = C"}\<close>
+ML_val \<open>const @{term "let A = (C::'a set) in A \<union> B"}\<close>
+ML_val \<open>const @{term "THE x::'b. P x"}\<close>
+ML_val \<open>const @{term "(\<lambda>x::'a. False)"}\<close>
+ML_val \<open>const @{term "(\<lambda>x::'a. True)"}\<close>
+ML_val \<open>const @{term "(\<lambda>x::'a. False) = (\<lambda>x::'a. False)"}\<close>
+ML_val \<open>const @{term "(\<lambda>x::'a. True) = (\<lambda>x::'a. True)"}\<close>
+ML_val \<open>const @{term "Let (a::'a) A"}\<close>
+ML_val \<open>const @{term "A (a::'a)"}\<close>
+ML_val \<open>const @{term "insert (a::'a) A = B"}\<close>
+ML_val \<open>const @{term "- (A::'a set)"}\<close>
+ML_val \<open>const @{term "finite (A::'a set)"}\<close>
+ML_val \<open>const @{term "\<not> finite (A::'a set)"}\<close>
+ML_val \<open>const @{term "finite (A::'a set set)"}\<close>
+ML_val \<open>const @{term "\<lambda>a::'a. A a \<and> \<not> B a"}\<close>
+ML_val \<open>const @{term "A < (B::'a set)"}\<close>
+ML_val \<open>const @{term "A \<le> (B::'a set)"}\<close>
+ML_val \<open>const @{term "[a::'a]"}\<close>
+ML_val \<open>const @{term "[a::'a set]"}\<close>
+ML_val \<open>const @{term "[A \<union> (B::'a set)]"}\<close>
+ML_val \<open>const @{term "[A \<union> (B::'a set)] = [C]"}\<close>
+ML_val \<open>const @{term "{(\<lambda>x::'a. x = a)} = C"}\<close>
+ML_val \<open>const @{term "(\<lambda>a::'a. \<not> A a) = B"}\<close>
+ML_val \<open>const @{prop "\<forall>F f g (h::'a set). F f \<and> F g \<and> \<not> f a \<and> g a \<longrightarrow> \<not> f a"}\<close>
+ML_val \<open>const @{term "\<lambda>A B x::'a. A x \<and> B x \<and> A = B"}\<close>
+ML_val \<open>const @{term "p = (\<lambda>(x::'a) (y::'a). P x \<or> \<not> Q y)"}\<close>
+ML_val \<open>const @{term "p = (\<lambda>(x::'a) (y::'a). p x y :: bool)"}\<close>
+ML_val \<open>const @{term "p = (\<lambda>A B x. A x \<and> \<not> B x) (\<lambda>x. True) (\<lambda>y. x \<noteq> y)"}\<close>
+ML_val \<open>const @{term "p = (\<lambda>y. x \<noteq> y)"}\<close>
+ML_val \<open>const @{term "(\<lambda>x. (p::'a\<Rightarrow>bool\<Rightarrow>bool) x False)"}\<close>
+ML_val \<open>const @{term "(\<lambda>x y. (p::'a\<Rightarrow>'a\<Rightarrow>bool\<Rightarrow>bool) x y False)"}\<close>
+ML_val \<open>const @{term "f = (\<lambda>x::'a. P x \<longrightarrow> Q x)"}\<close>
+ML_val \<open>const @{term "\<forall>a::'a. P a"}\<close>
-ML_val {* nonconst @{term "\<forall>P (a::'a). P a"} *}
-ML_val {* nonconst @{term "THE x::'a. P x"} *}
-ML_val {* nonconst @{term "SOME x::'a. P x"} *}
-ML_val {* nonconst @{term "(\<lambda>A B x::'a. A x \<or> B x) = myunion"} *}
-ML_val {* nonconst @{term "(\<lambda>x::'a. False) = (\<lambda>x::'a. True)"} *}
-ML_val {* nonconst @{prop "\<forall>F f g (h::'a set). F f \<and> F g \<and> \<not> a \<in> f \<and> a \<in> g \<longrightarrow> F h"} *}
+ML_val \<open>nonconst @{term "\<forall>P (a::'a). P a"}\<close>
+ML_val \<open>nonconst @{term "THE x::'a. P x"}\<close>
+ML_val \<open>nonconst @{term "SOME x::'a. P x"}\<close>
+ML_val \<open>nonconst @{term "(\<lambda>A B x::'a. A x \<or> B x) = myunion"}\<close>
+ML_val \<open>nonconst @{term "(\<lambda>x::'a. False) = (\<lambda>x::'a. True)"}\<close>
+ML_val \<open>nonconst @{prop "\<forall>F f g (h::'a set). F f \<and> F g \<and> \<not> a \<in> f \<and> a \<in> g \<longrightarrow> F h"}\<close>
-ML_val {* mono @{prop "Q (\<forall>x::'a set. P x)"} *}
-ML_val {* mono @{prop "P (a::'a)"} *}
-ML_val {* mono @{prop "{a} = {b::'a}"} *}
-ML_val {* mono @{prop "(\<lambda>x. x = a) = (\<lambda>y. y = (b::'a))"} *}
-ML_val {* mono @{prop "(a::'a) \<in> P \<and> P \<union> P = P"} *}
-ML_val {* mono @{prop "\<forall>F::'a set set. P"} *}
-ML_val {* mono @{prop "\<not> (\<forall>F f g (h::'a set). F f \<and> F g \<and> \<not> a \<in> f \<and> a \<in> g \<longrightarrow> F h)"} *}
-ML_val {* mono @{prop "\<not> Q (\<forall>x::'a set. P x)"} *}
-ML_val {* mono @{prop "\<not> (\<forall>x::'a. P x)"} *}
-ML_val {* mono @{prop "myall P = (P = (\<lambda>x::'a. True))"} *}
-ML_val {* mono @{prop "myall P = (P = (\<lambda>x::'a. False))"} *}
-ML_val {* mono @{prop "\<forall>x::'a. P x"} *}
-ML_val {* mono @{term "(\<lambda>A B x::'a. A x \<or> B x) \<noteq> myunion"} *}
+ML_val \<open>mono @{prop "Q (\<forall>x::'a set. P x)"}\<close>
+ML_val \<open>mono @{prop "P (a::'a)"}\<close>
+ML_val \<open>mono @{prop "{a} = {b::'a}"}\<close>
+ML_val \<open>mono @{prop "(\<lambda>x. x = a) = (\<lambda>y. y = (b::'a))"}\<close>
+ML_val \<open>mono @{prop "(a::'a) \<in> P \<and> P \<union> P = P"}\<close>
+ML_val \<open>mono @{prop "\<forall>F::'a set set. P"}\<close>
+ML_val \<open>mono @{prop "\<not> (\<forall>F f g (h::'a set). F f \<and> F g \<and> \<not> a \<in> f \<and> a \<in> g \<longrightarrow> F h)"}\<close>
+ML_val \<open>mono @{prop "\<not> Q (\<forall>x::'a set. P x)"}\<close>
+ML_val \<open>mono @{prop "\<not> (\<forall>x::'a. P x)"}\<close>
+ML_val \<open>mono @{prop "myall P = (P = (\<lambda>x::'a. True))"}\<close>
+ML_val \<open>mono @{prop "myall P = (P = (\<lambda>x::'a. False))"}\<close>
+ML_val \<open>mono @{prop "\<forall>x::'a. P x"}\<close>
+ML_val \<open>mono @{term "(\<lambda>A B x::'a. A x \<or> B x) \<noteq> myunion"}\<close>
-ML_val {* nonmono @{prop "A = (\<lambda>x::'a. True) \<and> A = (\<lambda>x. False)"} *}
-ML_val {* nonmono @{prop "\<forall>F f g (h::'a set). F f \<and> F g \<and> \<not> a \<in> f \<and> a \<in> g \<longrightarrow> F h"} *}
+ML_val \<open>nonmono @{prop "A = (\<lambda>x::'a. True) \<and> A = (\<lambda>x. False)"}\<close>
+ML_val \<open>nonmono @{prop "\<forall>F f g (h::'a set). F f \<and> F g \<and> \<not> a \<in> f \<and> a \<in> g \<longrightarrow> F h"}\<close>
-ML {*
+ML \<open>
val preproc_timeout = seconds 5.0
val mono_timeout = seconds 1.0
@@ -188,7 +188,7 @@
in File.append path (res ^ "\n"); writeln res end
handle Timeout.TIMEOUT _ => ()
in thy |> theorems_of |> List.app check_theorem end
-*}
+\<close>
(*
ML_val {* check_theory @{theory AVL2} *}
@@ -199,6 +199,6 @@
ML_val {* check_theory @{theory Relation} *}
*)
-ML {* getenv "ISABELLE_TMP" *}
+ML \<open>getenv "ISABELLE_TMP"\<close>
end
--- a/src/HOL/Nitpick_Examples/Pattern_Nits.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Nitpick_Examples/Pattern_Nits.thy Thu May 26 17:51:22 2016 +0200
@@ -5,7 +5,7 @@
Examples featuring Nitpick's "destroy_constrs" optimization.
*)
-section {* Examples Featuring Nitpick's \textit{destroy\_constrs} Optimization *}
+section \<open>Examples Featuring Nitpick's \textit{destroy\_constrs} Optimization\<close>
theory Pattern_Nits
imports Main
--- a/src/HOL/Nitpick_Examples/Record_Nits.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Nitpick_Examples/Record_Nits.thy Thu May 26 17:51:22 2016 +0200
@@ -5,7 +5,7 @@
Examples featuring Nitpick applied to records.
*)
-section {* Examples Featuring Nitpick Applied to Records *}
+section \<open>Examples Featuring Nitpick Applied to Records\<close>
theory Record_Nits
imports Main
--- a/src/HOL/Nitpick_Examples/Refute_Nits.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Nitpick_Examples/Refute_Nits.thy Thu May 26 17:51:22 2016 +0200
@@ -5,7 +5,7 @@
Refute examples adapted to Nitpick.
*)
-section {* Refute Examples Adapted to Nitpick *}
+section \<open>Refute Examples Adapted to Nitpick\<close>
theory Refute_Nits
imports Main
@@ -23,9 +23,9 @@
nitpick [sat_solver = SAT4J, expect = genuine] 2
oops
-subsection {* Examples and Test Cases *}
+subsection \<open>Examples and Test Cases\<close>
-subsubsection {* Propositional logic *}
+subsubsection \<open>Propositional logic\<close>
lemma "True"
nitpick [expect = none]
@@ -64,7 +64,7 @@
nitpick [expect = genuine]
oops
-subsubsection {* Predicate logic *}
+subsubsection \<open>Predicate logic\<close>
lemma "P x y z"
nitpick [expect = genuine]
@@ -78,7 +78,7 @@
nitpick [expect = genuine]
oops
-subsubsection {* Equality *}
+subsubsection \<open>Equality\<close>
lemma "P = True"
nitpick [expect = genuine]
@@ -110,7 +110,7 @@
nitpick [expect = genuine]
oops
-subsubsection {* First-Order Logic *}
+subsubsection \<open>First-Order Logic\<close>
lemma "\<exists>x. P x"
nitpick [expect = genuine]
@@ -148,14 +148,14 @@
nitpick [expect = genuine]
oops
-text {* A true statement (also testing names of free and bound variables being identical) *}
+text \<open>A true statement (also testing names of free and bound variables being identical)\<close>
lemma "(\<forall>x y. P x y \<longrightarrow> P y x) \<longrightarrow> (\<forall>x. P x y) \<longrightarrow> P y x"
nitpick [expect = none]
apply fast
done
-text {* "A type has at most 4 elements." *}
+text \<open>"A type has at most 4 elements."\<close>
lemma "\<not> distinct [a, b, c, d, e]"
nitpick [expect = genuine]
@@ -169,44 +169,44 @@
nitpick [expect = genuine]
oops
-text {* "Every reflexive and symmetric relation is transitive." *}
+text \<open>"Every reflexive and symmetric relation is transitive."\<close>
lemma "\<lbrakk>\<forall>x. P x x; \<forall>x y. P x y \<longrightarrow> P y x\<rbrakk> \<Longrightarrow> P x y \<longrightarrow> P y z \<longrightarrow> P x z"
nitpick [expect = genuine]
oops
-text {* The ``Drinker's theorem'' *}
+text \<open>The ``Drinker's theorem''\<close>
lemma "\<exists>x. f x = g x \<longrightarrow> f = g"
nitpick [expect = none]
apply (auto simp add: ext)
done
-text {* And an incorrect version of it *}
+text \<open>And an incorrect version of it\<close>
lemma "(\<exists>x. f x = g x) \<longrightarrow> f = g"
nitpick [expect = genuine]
oops
-text {* "Every function has a fixed point." *}
+text \<open>"Every function has a fixed point."\<close>
lemma "\<exists>x. f x = x"
nitpick [expect = genuine]
oops
-text {* "Function composition is commutative." *}
+text \<open>"Function composition is commutative."\<close>
lemma "f (g x) = g (f x)"
nitpick [expect = genuine]
oops
-text {* "Two functions that are equivalent wrt.\ the same predicate 'P' are equal." *}
+text \<open>"Two functions that are equivalent wrt.\ the same predicate 'P' are equal."\<close>
lemma "((P::('a\<Rightarrow>'b)\<Rightarrow>bool) f = P g) \<longrightarrow> (f x = g x)"
nitpick [expect = genuine]
oops
-subsubsection {* Higher-Order Logic *}
+subsubsection \<open>Higher-Order Logic\<close>
lemma "\<exists>P. P"
nitpick [expect = none]
@@ -242,7 +242,7 @@
nitpick [expect = genuine]
oops
-text {* ``The transitive closure of an arbitrary relation is non-empty.'' *}
+text \<open>``The transitive closure of an arbitrary relation is non-empty.''\<close>
definition "trans" :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool" where
"trans P \<equiv> (ALL x y z. P x y \<longrightarrow> P y z \<longrightarrow> P x z)"
@@ -259,7 +259,7 @@
nitpick [expect = genuine]
oops
-text {* ``The union of transitive closures is equal to the transitive closure of unions.'' *}
+text \<open>``The union of transitive closures is equal to the transitive closure of unions.''\<close>
lemma "(\<forall>x y. (P x y \<or> R x y) \<longrightarrow> T x y) \<longrightarrow> trans T \<longrightarrow> (\<forall>Q. (\<forall>x y. (P x y \<or> R x y) \<longrightarrow> Q x y) \<longrightarrow> trans Q \<longrightarrow> subset T Q)
\<longrightarrow> trans_closure TP P
@@ -268,19 +268,19 @@
nitpick [expect = genuine]
oops
-text {* ``Every surjective function is invertible.'' *}
+text \<open>``Every surjective function is invertible.''\<close>
lemma "(\<forall>y. \<exists>x. y = f x) \<longrightarrow> (\<exists>g. \<forall>x. g (f x) = x)"
nitpick [expect = genuine]
oops
-text {* ``Every invertible function is surjective.'' *}
+text \<open>``Every invertible function is surjective.''\<close>
lemma "(\<exists>g. \<forall>x. g (f x) = x) \<longrightarrow> (\<forall>y. \<exists>x. y = f x)"
nitpick [expect = genuine]
oops
-text {* ``Every point is a fixed point of some function.'' *}
+text \<open>``Every point is a fixed point of some function.''\<close>
lemma "\<exists>f. f x = x"
nitpick [card = 1-7, expect = none]
@@ -288,13 +288,13 @@
apply simp
done
-text {* Axiom of Choice: first an incorrect version *}
+text \<open>Axiom of Choice: first an incorrect version\<close>
lemma "(\<forall>x. \<exists>y. P x y) \<longrightarrow> (\<exists>!f. \<forall>x. P x (f x))"
nitpick [expect = genuine]
oops
-text {* And now two correct ones *}
+text \<open>And now two correct ones\<close>
lemma "(\<forall>x. \<exists>y. P x y) \<longrightarrow> (\<exists>f. \<forall>x. P x (f x))"
nitpick [card = 1-4, expect = none]
@@ -308,7 +308,7 @@
apply (fast intro: ext)
done
-subsubsection {* Metalogic *}
+subsubsection \<open>Metalogic\<close>
lemma "\<And>x. P x"
nitpick [expect = genuine]
@@ -338,7 +338,7 @@
nitpick [expect = genuine]
oops
-subsubsection {* Schematic Variables *}
+subsubsection \<open>Schematic Variables\<close>
schematic_goal "?P"
nitpick [expect = none]
@@ -350,7 +350,7 @@
apply auto
done
-subsubsection {* Abstractions *}
+subsubsection \<open>Abstractions\<close>
lemma "(\<lambda>x. x) = (\<lambda>x. y)"
nitpick [expect = genuine]
@@ -365,7 +365,7 @@
apply simp
done
-subsubsection {* Sets *}
+subsubsection \<open>Sets\<close>
lemma "P (A::'a set)"
nitpick [expect = genuine]
@@ -408,7 +408,7 @@
nitpick [expect = genuine]
oops
-subsubsection {* @{const undefined} *}
+subsubsection \<open>@{const undefined}\<close>
lemma "undefined"
nitpick [expect = genuine]
@@ -426,7 +426,7 @@
nitpick [expect = genuine]
oops
-subsubsection {* @{const The} *}
+subsubsection \<open>@{const The}\<close>
lemma "The P"
nitpick [expect = genuine]
@@ -448,7 +448,7 @@
nitpick [expect = genuine]
oops
-subsubsection {* @{const Eps} *}
+subsubsection \<open>@{const Eps}\<close>
lemma "Eps P"
nitpick [expect = genuine]
@@ -471,7 +471,7 @@
apply (auto simp add: someI)
done
-subsubsection {* Operations on Natural Numbers *}
+subsubsection \<open>Operations on Natural Numbers\<close>
lemma "(x::nat) + y = 0"
nitpick [expect = genuine]
@@ -493,7 +493,7 @@
nitpick [card = 1, expect = genuine]
oops
-text {* \<times> *}
+text \<open>\<times>\<close>
lemma "P (x::'a\<times>'b)"
nitpick [expect = genuine]
@@ -523,9 +523,9 @@
nitpick [expect = genuine]
oops
-subsubsection {* Subtypes (typedef), typedecl *}
+subsubsection \<open>Subtypes (typedef), typedecl\<close>
-text {* A completely unspecified non-empty subset of @{typ "'a"}: *}
+text \<open>A completely unspecified non-empty subset of @{typ "'a"}:\<close>
definition "myTdef = insert (undefined::'a) (undefined::'a set)"
@@ -547,9 +547,9 @@
nitpick [expect = genuine]
oops
-subsubsection {* Inductive Datatypes *}
+subsubsection \<open>Inductive Datatypes\<close>
-text {* unit *}
+text \<open>unit\<close>
lemma "P (x::unit)"
nitpick [expect = genuine]
@@ -567,7 +567,7 @@
nitpick [expect = genuine]
oops
-text {* option *}
+text \<open>option\<close>
lemma "P (x::'a option)"
nitpick [expect = genuine]
@@ -589,7 +589,7 @@
nitpick [expect = genuine]
oops
-text {* + *}
+text \<open>+\<close>
lemma "P (x::'a+'b)"
nitpick [expect = genuine]
@@ -615,7 +615,7 @@
nitpick [expect = genuine]
oops
-text {* Non-recursive datatypes *}
+text \<open>Non-recursive datatypes\<close>
datatype T1 = A | B
@@ -712,9 +712,9 @@
nitpick [expect = genuine]
oops
-text {* Recursive datatypes *}
+text \<open>Recursive datatypes\<close>
-text {* nat *}
+text \<open>nat\<close>
lemma "P (x::nat)"
nitpick [expect = genuine]
@@ -750,7 +750,7 @@
nitpick [expect = genuine]
oops
-text {* 'a list *}
+text \<open>'a list\<close>
lemma "P (xs::'a list)"
nitpick [expect = genuine]
@@ -855,7 +855,7 @@
nitpick [expect = genuine]
oops
-text {* Mutually recursive datatypes *}
+text \<open>Mutually recursive datatypes\<close>
datatype 'a aexp = Number 'a | ITE "'a bexp" "'a aexp" "'a aexp"
and 'a bexp = Equal "'a aexp" "'a aexp"
@@ -995,9 +995,9 @@
nitpick [expect = genuine]
oops
-text {* Other datatype examples *}
+text \<open>Other datatype examples\<close>
-text {* Indirect recursion is implemented via mutual recursion. *}
+text \<open>Indirect recursion is implemented via mutual recursion.\<close>
datatype XOpt = CX "XOpt option" | DX "bool \<Rightarrow> XOpt option"
@@ -1107,7 +1107,7 @@
nitpick [expect = genuine]
oops
-text {* Taken from "Inductive datatypes in HOL", p. 8: *}
+text \<open>Taken from "Inductive datatypes in HOL", p. 8:\<close>
datatype (dead 'a, 'b) T = C "'a \<Rightarrow> bool" | D "'b list"
datatype 'c U = E "('c, 'c U) T"
@@ -1124,7 +1124,7 @@
nitpick [expect = genuine]
oops
-subsubsection {* Records *}
+subsubsection \<open>Records\<close>
record ('a, 'b) point =
xpos :: 'a
@@ -1141,7 +1141,7 @@
nitpick [expect = genuine]
oops
-subsubsection {* Inductively Defined Sets *}
+subsubsection \<open>Inductively Defined Sets\<close>
inductive_set undefinedSet :: "'a set" where
"undefined \<in> undefinedSet"
@@ -1182,7 +1182,7 @@
nitpick [expect = genuine]
oops
-subsubsection {* Examples Involving Special Functions *}
+subsubsection \<open>Examples Involving Special Functions\<close>
lemma "card x = 0"
nitpick [expect = genuine]
@@ -1212,9 +1212,9 @@
nitpick [card = 2, expect = genuine]
oops
-subsubsection {* Axiomatic Type Classes and Overloading *}
+subsubsection \<open>Axiomatic Type Classes and Overloading\<close>
-text {* A type class without axioms: *}
+text \<open>A type class without axioms:\<close>
class classA
@@ -1222,7 +1222,7 @@
nitpick [expect = genuine]
oops
-text {* An axiom with a type variable (denoting types which have at least two elements): *}
+text \<open>An axiom with a type variable (denoting types which have at least two elements):\<close>
class classC =
assumes classC_ax: "\<exists>x y. x \<noteq> y"
@@ -1235,7 +1235,7 @@
nitpick [expect = none]
sorry
-text {* A type class for which a constant is defined: *}
+text \<open>A type class for which a constant is defined:\<close>
class classD =
fixes classD_const :: "'a \<Rightarrow> 'a"
@@ -1245,7 +1245,7 @@
nitpick [expect = genuine]
oops
-text {* A type class with multiple superclasses: *}
+text \<open>A type class with multiple superclasses:\<close>
class classE = classC + classD
@@ -1253,7 +1253,7 @@
nitpick [expect = genuine]
oops
-text {* OFCLASS: *}
+text \<open>OFCLASS:\<close>
lemma "OFCLASS('a::type, type_class)"
nitpick [expect = none]
@@ -1269,7 +1269,7 @@
nitpick [expect = genuine]
oops
-text {* Overloading: *}
+text \<open>Overloading:\<close>
consts inverse :: "'a \<Rightarrow> 'a"
--- a/src/HOL/Nitpick_Examples/Special_Nits.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Nitpick_Examples/Special_Nits.thy Thu May 26 17:51:22 2016 +0200
@@ -5,7 +5,7 @@
Examples featuring Nitpick's "specialize" optimization.
*)
-section {* Examples Featuring Nitpick's \textit{specialize} Optimization *}
+section \<open>Examples Featuring Nitpick's \textit{specialize} Optimization\<close>
theory Special_Nits
imports Main
--- a/src/HOL/Nitpick_Examples/Tests_Nits.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Nitpick_Examples/Tests_Nits.thy Thu May 26 17:51:22 2016 +0200
@@ -5,12 +5,12 @@
Nitpick tests.
*)
-section {* Nitpick Tests *}
+section \<open>Nitpick Tests\<close>
theory Tests_Nits
imports Main
begin
-ML {* () |> getenv "KODKODI" <> "" ? Nitpick_Tests.run_all_tests *}
+ML \<open>() |> getenv "KODKODI" <> "" ? Nitpick_Tests.run_all_tests\<close>
end
--- a/src/HOL/Nitpick_Examples/Typedef_Nits.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Nitpick_Examples/Typedef_Nits.thy Thu May 26 17:51:22 2016 +0200
@@ -5,7 +5,7 @@
Examples featuring Nitpick applied to typedefs.
*)
-section {* Examples Featuring Nitpick Applied to Typedefs *}
+section \<open>Examples Featuring Nitpick Applied to Typedefs\<close>
theory Typedef_Nits
imports Complex_Main
--- a/src/HOL/Nominal/Examples/CK_Machine.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Nominal/Examples/CK_Machine.thy Thu May 26 17:51:22 2016 +0200
@@ -2,7 +2,7 @@
imports "../Nominal"
begin
-text {*
+text \<open>
This theory establishes soundness and completeness for a CK-machine
with respect to a cbv-big-step semantics. The language includes
@@ -19,7 +19,7 @@
@{url "http://www.cs.indiana.edu/~rpjames/lm.pdf"}
@{url "http://www.cl.cam.ac.uk/teaching/2001/Semantics/"}
-*}
+\<close>
atom_decl name
@@ -37,7 +37,7 @@
| ZET "lam" (* zero test *)
| EQI "lam" "lam" (* equality test on numbers *)
-section {* Capture-Avoiding Substitution *}
+section \<open>Capture-Avoiding Substitution\<close>
nominal_primrec
subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam" ("_[_::=_]" [100,100,100] 100)
@@ -79,7 +79,7 @@
by (nominal_induct t avoiding: x y s rule: lam.strong_induct)
(auto simp add: calc_atm fresh_atm abs_fresh perm_nat_def)
-section {* Evaluation Contexts *}
+section \<open>Evaluation Contexts\<close>
datatype ctx =
Hole ("\<box>")
@@ -94,7 +94,7 @@
| CEQIL "ctx" "lam"
| CEQIR "lam" "ctx"
-text {* The operation of filling a term into a context: *}
+text \<open>The operation of filling a term into a context:\<close>
fun
filling :: "ctx \<Rightarrow> lam \<Rightarrow> lam" ("_\<lbrakk>_\<rbrakk>")
@@ -111,7 +111,7 @@
| "(CEQIL E t')\<lbrakk>t\<rbrakk> = EQI (E\<lbrakk>t\<rbrakk>) t'"
| "(CEQIR t' E)\<lbrakk>t\<rbrakk> = EQI t' (E\<lbrakk>t\<rbrakk>)"
-text {* The operation of composing two contexts: *}
+text \<open>The operation of composing two contexts:\<close>
fun
ctx_compose :: "ctx \<Rightarrow> ctx \<Rightarrow> ctx" ("_ \<circ> _")
@@ -132,7 +132,7 @@
shows "(E1 \<circ> E2)\<lbrakk>t\<rbrakk> = E1\<lbrakk>E2\<lbrakk>t\<rbrakk>\<rbrakk>"
by (induct E1 rule: ctx.induct) (auto)
-text {* Composing a list (stack) of contexts. *}
+text \<open>Composing a list (stack) of contexts.\<close>
fun
ctx_composes :: "ctx list \<Rightarrow> ctx" ("_\<down>")
@@ -140,7 +140,7 @@
"[]\<down> = \<box>"
| "(E#Es)\<down> = (Es\<down>) \<circ> E"
-section {* The CK-Machine *}
+section \<open>The CK-Machine\<close>
inductive
val :: "lam\<Rightarrow>bool"
@@ -192,7 +192,7 @@
shows "<e1,Es1> \<mapsto>* <e2,Es2>"
using a by (rule ms2) (rule ms1)
-section {* The Evaluation Relation (Big-Step Semantics) *}
+section \<open>The Evaluation Relation (Big-Step Semantics)\<close>
inductive
eval :: "lam\<Rightarrow>lam\<Rightarrow>bool" ("_ \<Down> _")
@@ -235,7 +235,7 @@
shows "t \<Down> t"
using a by (induct) (auto)
-text {* The Completeness Property: *}
+text \<open>The Completeness Property:\<close>
theorem eval_implies_machine_star_ctx:
assumes a: "t \<Down> t'"
@@ -249,7 +249,7 @@
shows "<t,[]> \<mapsto>* <t',[]>"
using a by (auto dest: eval_implies_machine_star_ctx)
-section {* The CBV Reduction Relation (Small-Step Semantics) *}
+section \<open>The CBV Reduction Relation (Small-Step Semantics)\<close>
lemma less_eqvt[eqvt]:
fixes pi::"name prm"
@@ -348,7 +348,7 @@
by (induct)
(auto simp add: eval_val cbv_star_eval dest: cbvs2)
-text {* The Soundness Property *}
+text \<open>The Soundness Property\<close>
theorem machine_star_implies_eval:
assumes a: "<t1,[]> \<mapsto>* <t2,[]>"
@@ -359,9 +359,9 @@
then show "t1 \<Down> t2" using b by (simp add: cbv_star_implies_eval)
qed
-section {* Typing *}
+section \<open>Typing\<close>
-text {* Types *}
+text \<open>Types\<close>
nominal_datatype ty =
tVAR "string"
@@ -378,18 +378,18 @@
by (induct T rule: ty.induct)
(auto simp add: fresh_string)
-text {* Typing Contexts *}
+text \<open>Typing Contexts\<close>
type_synonym tctx = "(name\<times>ty) list"
-text {* Sub-Typing Contexts *}
+text \<open>Sub-Typing Contexts\<close>
abbreviation
"sub_tctx" :: "tctx \<Rightarrow> tctx \<Rightarrow> bool" ("_ \<subseteq> _")
where
"\<Gamma>\<^sub>1 \<subseteq> \<Gamma>\<^sub>2 \<equiv> \<forall>x. x \<in> set \<Gamma>\<^sub>1 \<longrightarrow> x \<in> set \<Gamma>\<^sub>2"
-text {* Valid Typing Contexts *}
+text \<open>Valid Typing Contexts\<close>
inductive
valid :: "tctx \<Rightarrow> bool"
@@ -423,7 +423,7 @@
using a1 a2 a3
by (induct) (auto simp add: fresh_set fresh_prod fresh_atm)
-section {* The Typing Relation *}
+section \<open>The Typing Relation\<close>
inductive
typing :: "tctx \<Rightarrow> lam \<Rightarrow> ty \<Rightarrow> bool" ("_ \<turnstile> _ : _")
@@ -474,7 +474,7 @@
by (cases rule: typing.strong_cases)
(auto simp add: alpha lam.inject abs_fresh ty_fresh)
-section {* The Type-Preservation Property for the CBV Reduction Relation *}
+section \<open>The Type-Preservation Property for the CBV Reduction Relation\<close>
lemma weakening:
fixes \<Gamma>1 \<Gamma>2::"tctx"
@@ -534,7 +534,7 @@
using a b
by (induct) (auto intro: cbv_type_preservation)
-section {* The Type-Preservation Property for the Machine and Evaluation Relation *}
+section \<open>The Type-Preservation Property for the Machine and Evaluation Relation\<close>
theorem machine_type_preservation:
assumes a: "<t,[]> \<mapsto>* <t',[]>"
@@ -554,7 +554,7 @@
then show "\<Gamma> \<turnstile> t' : T" using b by (simp add: machine_type_preservation)
qed
-text {* The Progress Property *}
+text \<open>The Progress Property\<close>
lemma canonical_tARR[dest]:
assumes a: "[] \<turnstile> t : T1 \<rightarrow> T2"
--- a/src/HOL/Nominal/Examples/CR.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Nominal/Examples/CR.thy Thu May 26 17:51:22 2016 +0200
@@ -2,7 +2,7 @@
imports Lam_Funs
begin
-text {* The Church-Rosser proof from Barendregt's book *}
+text \<open>The Church-Rosser proof from Barendregt's book\<close>
lemma forget:
assumes asm: "x\<sharp>L"
@@ -95,23 +95,23 @@
proof -
{ (*Case 1.1*)
assume "z=x"
- have "(1)": "?LHS = N[y::=L]" using `z=x` by simp
- have "(2)": "?RHS = N[y::=L]" using `z=x` `x\<noteq>y` by simp
+ have "(1)": "?LHS = N[y::=L]" using \<open>z=x\<close> by simp
+ have "(2)": "?RHS = N[y::=L]" using \<open>z=x\<close> \<open>x\<noteq>y\<close> by simp
from "(1)" "(2)" have "?LHS = ?RHS" by simp
}
moreover
{ (*Case 1.2*)
assume "z=y" and "z\<noteq>x"
- have "(1)": "?LHS = L" using `z\<noteq>x` `z=y` by simp
- have "(2)": "?RHS = L[x::=N[y::=L]]" using `z=y` by simp
- have "(3)": "L[x::=N[y::=L]] = L" using `x\<sharp>L` by (simp add: forget)
+ have "(1)": "?LHS = L" using \<open>z\<noteq>x\<close> \<open>z=y\<close> by simp
+ have "(2)": "?RHS = L[x::=N[y::=L]]" using \<open>z=y\<close> by simp
+ have "(3)": "L[x::=N[y::=L]] = L" using \<open>x\<sharp>L\<close> by (simp add: forget)
from "(1)" "(2)" "(3)" have "?LHS = ?RHS" by simp
}
moreover
{ (*Case 1.3*)
assume "z\<noteq>x" and "z\<noteq>y"
- have "(1)": "?LHS = Var z" using `z\<noteq>x` `z\<noteq>y` by simp
- have "(2)": "?RHS = Var z" using `z\<noteq>x` `z\<noteq>y` by simp
+ have "(1)": "?LHS = Var z" using \<open>z\<noteq>x\<close> \<open>z\<noteq>y\<close> by simp
+ have "(2)": "?RHS = Var z" using \<open>z\<noteq>x\<close> \<open>z\<noteq>y\<close> by simp
from "(1)" "(2)" have "?LHS = ?RHS" by simp
}
ultimately show "?LHS = ?RHS" by blast
@@ -125,10 +125,10 @@
hence "z\<sharp>N[y::=L]" by (simp add: fresh_fact)
show "(Lam [z].M1)[x::=N][y::=L] = (Lam [z].M1)[y::=L][x::=N[y::=L]]" (is "?LHS=?RHS")
proof -
- have "?LHS = Lam [z].(M1[x::=N][y::=L])" using `z\<sharp>x` `z\<sharp>y` `z\<sharp>N` `z\<sharp>L` by simp
- also from ih have "\<dots> = Lam [z].(M1[y::=L][x::=N[y::=L]])" using `x\<noteq>y` `x\<sharp>L` by simp
- also have "\<dots> = (Lam [z].(M1[y::=L]))[x::=N[y::=L]]" using `z\<sharp>x` `z\<sharp>N[y::=L]` by simp
- also have "\<dots> = ?RHS" using `z\<sharp>y` `z\<sharp>L` by simp
+ have "?LHS = Lam [z].(M1[x::=N][y::=L])" using \<open>z\<sharp>x\<close> \<open>z\<sharp>y\<close> \<open>z\<sharp>N\<close> \<open>z\<sharp>L\<close> by simp
+ also from ih have "\<dots> = Lam [z].(M1[y::=L][x::=N[y::=L]])" using \<open>x\<noteq>y\<close> \<open>x\<sharp>L\<close> by simp
+ also have "\<dots> = (Lam [z].(M1[y::=L]))[x::=N[y::=L]]" using \<open>z\<sharp>x\<close> \<open>z\<sharp>N[y::=L]\<close> by simp
+ also have "\<dots> = ?RHS" using \<open>z\<sharp>y\<close> \<open>z\<sharp>L\<close> by simp
finally show "?LHS = ?RHS" .
qed
next
@@ -143,7 +143,7 @@
by (nominal_induct M avoiding: x y N L rule: lam.strong_induct)
(auto simp add: fresh_fact forget)
-section {* Beta Reduction *}
+section \<open>Beta Reduction\<close>
inductive
"Beta" :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^sub>\<beta> _" [80,80] 80)
@@ -173,7 +173,7 @@
using a2 a1
by (induct) (auto)
-section {* One-Reduction *}
+section \<open>One-Reduction\<close>
inductive
One :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^sub>1 _" [80,80] 80)
@@ -287,7 +287,7 @@
by (cases rule: One.strong_cases [where a="a" and aa="a"])
(auto dest: one_abs simp add: lam.inject abs_fresh alpha fresh_prod)
-text {* first case in Lemma 3.2.4*}
+text \<open>first case in Lemma 3.2.4\<close>
lemma one_subst_aux:
assumes a: "N\<longrightarrow>\<^sub>1N'"
--- a/src/HOL/Nominal/Examples/CR_Takahashi.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Nominal/Examples/CR_Takahashi.thy Thu May 26 17:51:22 2016 +0200
@@ -30,7 +30,7 @@
apply(fresh_guess)+
done
-section {* Lemmas about Capture-Avoiding Substitution *}
+section \<open>Lemmas about Capture-Avoiding Substitution\<close>
lemma subst_eqvt[eqvt]:
fixes pi::"name prm"
@@ -63,7 +63,7 @@
by (nominal_induct t avoiding: x y s rule: lam.strong_induct)
(auto simp add: swap_simps fresh_atm abs_fresh)
-section {* Beta-Reduction *}
+section \<open>Beta-Reduction\<close>
inductive
"Beta" :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^sub>\<beta> _" [80,80] 80)
@@ -73,7 +73,7 @@
| b3[intro]: "t1 \<longrightarrow>\<^sub>\<beta> t2 \<Longrightarrow> Lam [x].t1 \<longrightarrow>\<^sub>\<beta> Lam [x].t2"
| b4[intro]: "App (Lam [x].t) s \<longrightarrow>\<^sub>\<beta> t[x::=s]"
-section {* Transitive Closure of Beta *}
+section \<open>Transitive Closure of Beta\<close>
inductive
"Beta_star" :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^sub>\<beta>\<^sup>* _" [80,80] 80)
@@ -82,7 +82,7 @@
| bs2[intro]: "t \<longrightarrow>\<^sub>\<beta> s \<Longrightarrow> t \<longrightarrow>\<^sub>\<beta>\<^sup>* s"
| bs3[intro,trans]: "\<lbrakk>t1\<longrightarrow>\<^sub>\<beta>\<^sup>* t2; t2 \<longrightarrow>\<^sub>\<beta>\<^sup>* t3\<rbrakk> \<Longrightarrow> t1 \<longrightarrow>\<^sub>\<beta>\<^sup>* t3"
-section {* One-Reduction *}
+section \<open>One-Reduction\<close>
inductive
One :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^sub>1 _" [80,80] 80)
@@ -145,7 +145,7 @@
by (cases rule: One.strong_cases)
(auto dest: One_Lam simp add: lam.inject abs_fresh alpha fresh_prod)
-section {* Transitive Closure of One *}
+section \<open>Transitive Closure of One\<close>
inductive
"One_star" :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^sub>1\<^sup>* _" [80,80] 80)
@@ -154,7 +154,7 @@
| os2[intro]: "t \<longrightarrow>\<^sub>1 s \<Longrightarrow> t \<longrightarrow>\<^sub>1\<^sup>* s"
| os3[intro]: "\<lbrakk>t1\<longrightarrow>\<^sub>1\<^sup>* t2; t2 \<longrightarrow>\<^sub>1\<^sup>* t3\<rbrakk> \<Longrightarrow> t1 \<longrightarrow>\<^sub>1\<^sup>* t3"
-section {* Complete Development Reduction *}
+section \<open>Complete Development Reduction\<close>
inductive
Dev :: "lam \<Rightarrow> lam \<Rightarrow> bool" (" _ \<longrightarrow>\<^sub>d _" [80,80]80)
@@ -248,7 +248,7 @@
shows "\<exists>t3. t2 \<longrightarrow>\<^sub>1\<^sup>* t3 \<and> t1 \<longrightarrow>\<^sub>1\<^sup>* t3"
using a Rectangle_for_One by (induct arbitrary: t2) (blast)+
-section {* Establishing the Equivalence of Beta-star and One-star *}
+section \<open>Establishing the Equivalence of Beta-star and One-star\<close>
lemma Beta_Lam_cong:
assumes a: "t1 \<longrightarrow>\<^sub>\<beta>\<^sup>* t2"
@@ -295,7 +295,7 @@
then show "t1 \<longrightarrow>\<^sub>1\<^sup>* t2" by (induct) (auto intro: Beta_implies_One_star)
qed
-section {* The Church-Rosser Theorem *}
+section \<open>The Church-Rosser Theorem\<close>
theorem CR_for_Beta_star:
assumes a: "t \<longrightarrow>\<^sub>\<beta>\<^sup>* t1" "t\<longrightarrow>\<^sub>\<beta>\<^sup>* t2"
--- a/src/HOL/Nominal/Examples/Class1.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Nominal/Examples/Class1.thy Thu May 26 17:51:22 2016 +0200
@@ -2,11 +2,11 @@
imports "../Nominal"
begin
-section {* Term-Calculus from Urban's PhD *}
+section \<open>Term-Calculus from Urban's PhD\<close>
atom_decl name coname
-text {* types *}
+text \<open>types\<close>
nominal_datatype ty =
PR "string"
@@ -44,7 +44,7 @@
by (nominal_induct T rule: ty.strong_induct)
(auto simp add: fresh_string)
-text {* terms *}
+text \<open>terms\<close>
nominal_datatype trm =
Ax "name" "coname"
@@ -60,15 +60,15 @@
| ImpR "\<guillemotleft>name\<guillemotright>(\<guillemotleft>coname\<guillemotright>trm)" "coname" ("ImpR (_).<_>._ _" [100,100,100,100] 100)
| ImpL "\<guillemotleft>coname\<guillemotright>trm" "\<guillemotleft>name\<guillemotright>trm" "name" ("ImpL <_>._ (_)._ _" [100,100,100,100,100] 100)
-text {* named terms *}
+text \<open>named terms\<close>
nominal_datatype ntrm = Na "\<guillemotleft>name\<guillemotright>trm" ("((_):_)" [100,100] 100)
-text {* conamed terms *}
+text \<open>conamed terms\<close>
nominal_datatype ctrm = Co "\<guillemotleft>coname\<guillemotright>trm" ("(<_>:_)" [100,100] 100)
-text {* renaming functions *}
+text \<open>renaming functions\<close>
nominal_primrec (freshness_context: "(d::coname,e::coname)")
crename :: "trm \<Rightarrow> coname \<Rightarrow> coname \<Rightarrow> trm" ("_[_\<turnstile>c>_]" [100,100,100] 100)
@@ -307,7 +307,7 @@
apply(simp_all add: trm.inject split: if_splits)
done
-text {* substitution functions *}
+text \<open>substitution functions\<close>
lemma fresh_perm_coname:
fixes c::"coname"
@@ -2835,7 +2835,7 @@
lemmas subst_comm' = substn_crename_comm' substc_crename_comm'
substn_nrename_comm' substc_nrename_comm'
-text {* typing contexts *}
+text \<open>typing contexts\<close>
type_synonym ctxtn = "(name\<times>ty) list"
type_synonym ctxtc = "(coname\<times>ty) list"
@@ -2868,7 +2868,7 @@
show "x\<sharp>\<Delta>" by (induct \<Delta>) (auto simp add: fresh_list_nil fresh_list_cons fresh_prod fresh_atm fresh_ty)
qed
-text {* cut-reductions *}
+text \<open>cut-reductions\<close>
declare abs_perm[eqvt]
@@ -4692,7 +4692,7 @@
finally show ?thesis by simp
qed
-text {* axioms do not reduce *}
+text \<open>axioms do not reduce\<close>
lemma ax_do_not_l_reduce:
shows "Ax x a \<longrightarrow>\<^sub>l M \<Longrightarrow> False"
@@ -5160,7 +5160,7 @@
apply(simp add: calc_atm perm_swap)
done
-text {* Transitive Closure*}
+text \<open>Transitive Closure\<close>
abbreviation
a_star_redu :: "trm \<Rightarrow> trm \<Rightarrow> bool" ("_ \<longrightarrow>\<^sub>a* _" [100,100] 100)
@@ -5189,7 +5189,7 @@
using a2 a1
by (induct) (auto)
-text {* congruence rules for \<open>\<longrightarrow>\<^sub>a*\<close> *}
+text \<open>congruence rules for \<open>\<longrightarrow>\<^sub>a*\<close>\<close>
lemma ax_do_not_a_star_reduce:
shows "Ax x a \<longrightarrow>\<^sub>a* M \<Longrightarrow> M = Ax x a"
@@ -5379,7 +5379,7 @@
apply(auto simp add: alpha trm.inject)
done
-text {* Substitution *}
+text \<open>Substitution\<close>
lemma subst_not_fin1:
shows "\<not>fin(M{x:=<c>.P}) x"
@@ -5644,7 +5644,7 @@
apply(drule fic_elims, simp)
done
-text {* Reductions *}
+text \<open>Reductions\<close>
lemma fin_l_reduce:
assumes a: "fin M x"
@@ -5782,7 +5782,7 @@
apply(auto simp add: fic_a_reduce)
done
-text {* substitution properties *}
+text \<open>substitution properties\<close>
lemma subst_with_ax1:
shows "M{x:=<a>.Ax y a} \<longrightarrow>\<^sub>a* M[x\<turnstile>n>y]"
@@ -6295,7 +6295,7 @@
using fs by simp
qed
-text {* substitution lemmas *}
+text \<open>substitution lemmas\<close>
lemma not_Ax1:
shows "\<not>(b\<sharp>M) \<Longrightarrow> M{b:=(y).Q} \<noteq> Ax x a"
--- a/src/HOL/Nominal/Examples/Class2.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Nominal/Examples/Class2.thy Thu May 26 17:51:22 2016 +0200
@@ -2,7 +2,7 @@
imports Class1
begin
-text {* Reduction *}
+text \<open>Reduction\<close>
lemma fin_not_Cut:
assumes a: "fin M x"
@@ -1821,9 +1821,9 @@
apply(auto)
done
-text {* Candidates and SN *}
+text \<open>Candidates and SN\<close>
-text {* SNa *}
+text \<open>SNa\<close>
inductive
SNa :: "trm \<Rightarrow> bool"
@@ -2091,7 +2091,7 @@
apply(perm_simp)
done
-text {* set operators *}
+text \<open>set operators\<close>
definition AXIOMSn :: "ty \<Rightarrow> ntrm set" where
"AXIOMSn B \<equiv> { (x):(Ax y b) | x y b. True }"
@@ -2873,7 +2873,7 @@
apply(simp_all)
done
-text {* Candidates *}
+text \<open>Candidates\<close>
lemma test1:
shows "x\<in>(X\<union>Y) = (x\<in>X \<or> x\<in>Y)"
@@ -3447,7 +3447,7 @@
}
qed
-text {* Elimination rules for the set-operators *}
+text \<open>Elimination rules for the set-operators\<close>
lemma BINDINGc_elim:
assumes a: "<a>:M \<in> BINDINGc B (\<parallel>(B)\<parallel>)"
@@ -5178,7 +5178,7 @@
apply(perm_simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt)
done
-text {* Main lemma 1 *}
+text \<open>Main lemma 1\<close>
lemma AXIOMS_imply_SNa:
shows "<a>:M \<in> AXIOMSc B \<Longrightarrow> SNa M"
@@ -5484,7 +5484,7 @@
}
qed
-text {* Main lemma 2 *}
+text \<open>Main lemma 2\<close>
lemma AXIOMS_preserved:
shows "<a>:M \<in> AXIOMSc B \<Longrightarrow> M \<longrightarrow>\<^sub>a* M' \<Longrightarrow> <a>:M' \<in> AXIOMSc B"
--- a/src/HOL/Nominal/Examples/Class3.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Nominal/Examples/Class3.thy Thu May 26 17:51:22 2016 +0200
@@ -2,7 +2,7 @@
imports Class2
begin
-text {* 3rd Main Lemma *}
+text \<open>3rd Main Lemma\<close>
lemma Cut_a_redu_elim:
assumes a: "Cut <a>.M (x).N \<longrightarrow>\<^sub>a R"
@@ -2246,7 +2246,7 @@
apply(blast)+
done
-text {* helper-stuff to set up the induction *}
+text \<open>helper-stuff to set up the induction\<close>
abbreviation
SNa_set :: "trm set"
@@ -2359,7 +2359,7 @@
apply(simp add: alpha fresh_prod fresh_atm)
done
-text {* 3rd lemma *}
+text \<open>3rd lemma\<close>
lemma CUT_SNa_aux:
assumes a1: "<a>:M \<in> (\<parallel><B>\<parallel>)"
@@ -4927,7 +4927,7 @@
done
-text {* typing relation *}
+text \<open>typing relation\<close>
inductive
typing :: "ctxtn \<Rightarrow> trm \<Rightarrow> ctxtc \<Rightarrow> bool" ("_ \<turnstile> _ \<turnstile> _" [100,100,100] 100)
--- a/src/HOL/Nominal/Examples/Compile.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Nominal/Examples/Compile.thy Thu May 26 17:51:22 2016 +0200
@@ -46,7 +46,7 @@
| ISeq "trmI" "trmI" ("_;;_" [100,100] 100)
| Iif "trmI" "trmI" "trmI"
-text {* valid contexts *}
+text \<open>valid contexts\<close>
inductive
valid :: "(name\<times>'a::pt_name) list \<Rightarrow> bool"
@@ -54,7 +54,7 @@
v1[intro]: "valid []"
| v2[intro]: "\<lbrakk>valid \<Gamma>;a\<sharp>\<Gamma>\<rbrakk>\<Longrightarrow> valid ((a,\<sigma>)#\<Gamma>)" (* maybe dom of \<Gamma> *)
-text {* typing judgements for trms *}
+text \<open>typing judgements for trms\<close>
inductive
typing :: "(name\<times>ty) list\<Rightarrow>trm\<Rightarrow>ty\<Rightarrow>bool" (" _ \<turnstile> _ : _ " [80,80,80] 80)
@@ -72,7 +72,7 @@
((x1,Data(\<sigma>1))#\<Gamma>) \<turnstile> e1 : \<tau>; ((x2,Data(\<sigma>2))#\<Gamma>) \<turnstile> e2 : \<tau>\<rbrakk>
\<Longrightarrow> \<Gamma> \<turnstile> (Case e of inl x1 \<rightarrow> e1 | inr x2 \<rightarrow> e2) : \<tau>"
-text {* typing judgements for Itrms *}
+text \<open>typing judgements for Itrms\<close>
inductive
Ityping :: "(name\<times>tyI) list\<Rightarrow>trmI\<Rightarrow>tyI\<Rightarrow>bool" (" _ I\<turnstile> _ : _ " [80,80,80] 80)
@@ -88,7 +88,7 @@
| t8[intro]: "\<lbrakk>\<Gamma> I\<turnstile> e1 : DataI(NatI); \<Gamma> I\<turnstile> e2 : \<tau>\<rbrakk> \<Longrightarrow> \<Gamma> I\<turnstile> e1;;e2 : \<tau>"
| t9[intro]: "\<lbrakk>\<Gamma> I\<turnstile> e: DataI(NatI); \<Gamma> I\<turnstile> e1 : \<tau>; \<Gamma> I\<turnstile> e2 : \<tau>\<rbrakk> \<Longrightarrow> \<Gamma> I\<turnstile> Iif e e1 e2 : \<tau>"
-text {* capture-avoiding substitution *}
+text \<open>capture-avoiding substitution\<close>
class subst =
fixes subst :: "'a \<Rightarrow> name \<Rightarrow> 'a \<Rightarrow> 'a" ("_[_::=_]" [100,100,100] 100)
@@ -179,7 +179,7 @@
apply(simp add: abs_supp)
done
-text {* big-step evaluation for trms *}
+text \<open>big-step evaluation for trms\<close>
inductive
big :: "trm\<Rightarrow>trm\<Rightarrow>bool" ("_ \<Down> _" [80,80] 80)
@@ -210,7 +210,7 @@
| m8[intro]: "\<lbrakk>(m,e)I\<Down>(m',INat(n)); n\<noteq>0; (m',e1)I\<Down>(m'',e)\<rbrakk> \<Longrightarrow> (m,Iif e e1 e2)I\<Down>(m'',e)"
| m9[intro]: "\<lbrakk>(m,e)I\<Down>(m',INat(0)); (m',e2)I\<Down>(m'',e)\<rbrakk> \<Longrightarrow> (m,Iif e e1 e2)I\<Down>(m'',e)"
-text {* Translation functions *}
+text \<open>Translation functions\<close>
nominal_primrec
trans :: "trm \<Rightarrow> trmI"
--- a/src/HOL/Nominal/Examples/Contexts.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Nominal/Examples/Contexts.thy Thu May 26 17:51:22 2016 +0200
@@ -2,7 +2,7 @@
imports "../Nominal"
begin
-text {*
+text \<open>
We show here that the Plotkin-style of defining
beta-reduction (based on congruence rules) is
@@ -14,23 +14,23 @@
of filling a term into a context produces an alpha-equivalent
term.
-*}
+\<close>
atom_decl name
-text {*
+text \<open>
Lambda-Terms - the Lam-term constructor binds a name
-*}
+\<close>
nominal_datatype lam =
Var "name"
| App "lam" "lam"
| Lam "\<guillemotleft>name\<guillemotright>lam" ("Lam [_]._" [100,100] 100)
-text {*
+text \<open>
Contexts - the lambda case in contexts does not bind
a name, even if we introduce the notation [_]._ for CLam.
-*}
+\<close>
nominal_datatype ctx =
Hole ("\<box>" 1000)
@@ -38,7 +38,7 @@
| CAppR "lam" "ctx"
| CLam "name" "ctx" ("CLam [_]._" [100,100] 100)
-text {* Capture-Avoiding Substitution *}
+text \<open>Capture-Avoiding Substitution\<close>
nominal_primrec
subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam" ("_[_::=_]" [100,100,100] 100)
@@ -52,10 +52,10 @@
apply(fresh_guess)+
done
-text {*
+text \<open>
Filling is the operation that fills a term into a hole.
This operation is possibly capturing.
-*}
+\<close>
nominal_primrec
filling :: "ctx \<Rightarrow> lam \<Rightarrow> lam" ("_\<lbrakk>_\<rbrakk>" [100,100] 100)
@@ -66,17 +66,17 @@
| "(CLam [x].E)\<lbrakk>t\<rbrakk> = Lam [x].(E\<lbrakk>t\<rbrakk>)"
by (rule TrueI)+
-text {*
+text \<open>
While contexts themselves are not alpha-equivalence classes, the
filling operation produces an alpha-equivalent lambda-term.
-*}
+\<close>
lemma alpha_test:
shows "x\<noteq>y \<Longrightarrow> (CLam [x].\<box>) \<noteq> (CLam [y].\<box>)"
and "(CLam [x].\<box>)\<lbrakk>Var x\<rbrakk> = (CLam [y].\<box>)\<lbrakk>Var y\<rbrakk>"
by (auto simp add: alpha ctx.inject lam.inject calc_atm fresh_atm)
-text {* The composition of two contexts. *}
+text \<open>The composition of two contexts.\<close>
nominal_primrec
ctx_compose :: "ctx \<Rightarrow> ctx \<Rightarrow> ctx" ("_ \<circ> _" [100,100] 100)
@@ -91,14 +91,14 @@
shows "(E1 \<circ> E2)\<lbrakk>t\<rbrakk> = E1\<lbrakk>E2\<lbrakk>t\<rbrakk>\<rbrakk>"
by (induct E1 rule: ctx.induct) (auto)
-text {* Beta-reduction via contexts in the Felleisen-Hieb style. *}
+text \<open>Beta-reduction via contexts in the Felleisen-Hieb style.\<close>
inductive
ctx_red :: "lam\<Rightarrow>lam\<Rightarrow>bool" ("_ \<longrightarrow>x _" [80,80] 80)
where
xbeta[intro]: "E\<lbrakk>App (Lam [x].t) t'\<rbrakk> \<longrightarrow>x E\<lbrakk>t[x::=t']\<rbrakk>"
-text {* Beta-reduction via congruence rules in the Plotkin style. *}
+text \<open>Beta-reduction via congruence rules in the Plotkin style.\<close>
inductive
cong_red :: "lam\<Rightarrow>lam\<Rightarrow>bool" ("_ \<longrightarrow>o _" [80,80] 80)
@@ -108,7 +108,7 @@
| oapp2[intro]: "t \<longrightarrow>o t' \<Longrightarrow> App t2 t \<longrightarrow>o App t2 t'"
| olam[intro]: "t \<longrightarrow>o t' \<Longrightarrow> Lam [x].t \<longrightarrow>o Lam [x].t'"
-text {* The proof that shows both relations are equal. *}
+text \<open>The proof that shows both relations are equal.\<close>
lemma cong_red_in_ctx:
assumes a: "t \<longrightarrow>o t'"
--- a/src/HOL/Nominal/Examples/Crary.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Nominal/Examples/Crary.thy Thu May 26 17:51:22 2016 +0200
@@ -65,7 +65,7 @@
shows "size T > 0"
by (nominal_induct rule: ty.strong_induct) (simp_all)
-section {* Substitutions *}
+section \<open>Substitutions\<close>
fun
lookup :: "Subst \<Rightarrow> name \<Rightarrow> trm"
@@ -259,7 +259,7 @@
then show "\<theta><(Lam [n].t)[x::=u]> = \<theta><Lam [n].t>[x::=\<theta><u>]" using fs by auto
qed (auto)
-section {* Typing *}
+section \<open>Typing\<close>
inductive
valid :: "Ctxt \<Rightarrow> bool"
@@ -336,7 +336,7 @@
declare ty.inject [simp del]
-section {* Definitional Equivalence *}
+section \<open>Definitional Equivalence\<close>
inductive
def_equiv :: "Ctxt\<Rightarrow>trm\<Rightarrow>trm\<Rightarrow>ty\<Rightarrow>bool" ("_ \<turnstile> _ \<equiv> _ : _" [60,60] 60)
@@ -362,7 +362,7 @@
shows "valid \<Gamma>"
using a by (induct) (auto elim: typing_implies_valid)
-section {* Weak Head Reduction *}
+section \<open>Weak Head Reduction\<close>
inductive
whr_def :: "trm\<Rightarrow>trm\<Rightarrow>bool" ("_ \<leadsto> _" [80,80] 80)
@@ -389,7 +389,7 @@
equivariance whr_def
-section {* Weak Head Normalisation *}
+section \<open>Weak Head Normalisation\<close>
abbreviation
nf :: "trm \<Rightarrow> bool" ("_ \<leadsto>|" [100] 100)
@@ -440,7 +440,7 @@
qed (auto)
-section {* Algorithmic Term Equivalence and Algorithmic Path Equivalence *}
+section \<open>Algorithmic Term Equivalence and Algorithmic Path Equivalence\<close>
inductive
alg_equiv :: "Ctxt\<Rightarrow>trm\<Rightarrow>trm\<Rightarrow>ty\<Rightarrow>bool" ("_ \<turnstile> _ \<Leftrightarrow> _ : _" [60,60,60,60] 60)
@@ -605,7 +605,7 @@
using assms
by (induct rule: alg_equiv_alg_path_equiv.inducts(2)) (simp, auto)
-section {* Logical Equivalence *}
+section \<open>Logical Equivalence\<close>
function log_equiv :: "(Ctxt \<Rightarrow> trm \<Rightarrow> trm \<Rightarrow> ty \<Rightarrow> bool)" ("_ \<turnstile> _ is _ : _" [60,60,60,60] 60)
where
@@ -955,10 +955,10 @@
then show "\<Gamma> \<turnstile> s \<Leftrightarrow> t : T" using main_lemma(1) val by simp
qed
-text {* We leave soundness as an exercise - just like Crary in the ATS book :-) \\
+text \<open>We leave soundness as an exercise - just like Crary in the ATS book :-) \\
@{prop[mode=IfThen] "\<lbrakk>\<Gamma> \<turnstile> s \<Leftrightarrow> t : T; \<Gamma> \<turnstile> t : T; \<Gamma> \<turnstile> s : T\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s \<equiv> t : T"} \\
@{prop "\<lbrakk>\<Gamma> \<turnstile> s \<leftrightarrow> t : T; \<Gamma> \<turnstile> t : T; \<Gamma> \<turnstile> s : T\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s \<equiv> t : T"}
-*}
+\<close>
end
--- a/src/HOL/Nominal/Examples/Fsub.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Nominal/Examples/Fsub.thy Thu May 26 17:51:22 2016 +0200
@@ -4,7 +4,7 @@
begin
(*>*)
-text{* Authors: Christian Urban,
+text\<open>Authors: Christian Urban,
Benjamin Pierce,
Dimitrios Vytiniotis
Stephanie Weirich
@@ -12,26 +12,26 @@
Julien Narboux
Stefan Berghofer
- with great help from Markus Wenzel. *}
+ with great help from Markus Wenzel.\<close>
-section {* Types for Names, Nominal Datatype Declaration for Types and Terms *}
+section \<open>Types for Names, Nominal Datatype Declaration for Types and Terms\<close>
no_syntax
"_Map" :: "maplets => 'a \<rightharpoonup> 'b" ("(1[_])")
-text {* The main point of this solution is to use names everywhere (be they bound,
+text \<open>The main point of this solution is to use names everywhere (be they bound,
binding or free). In System \FSUB{} there are two kinds of names corresponding to
type-variables and to term-variables. These two kinds of names are represented in
- the nominal datatype package as atom-types @{text "tyvrs"} and @{text "vrs"}: *}
+ the nominal datatype package as atom-types \<open>tyvrs\<close> and \<open>vrs\<close>:\<close>
atom_decl tyvrs vrs
-text{* There are numerous facts that come with this declaration: for example that
- there are infinitely many elements in @{text "tyvrs"} and @{text "vrs"}. *}
+text\<open>There are numerous facts that come with this declaration: for example that
+ there are infinitely many elements in \<open>tyvrs\<close> and \<open>vrs\<close>.\<close>
-text{* The constructors for types and terms in System \FSUB{} contain abstractions
+text\<open>The constructors for types and terms in System \FSUB{} contain abstractions
over type-variables and term-variables. The nominal datatype package uses
- @{text "\<guillemotleft>\<dots>\<guillemotright>\<dots>"} to indicate where abstractions occur. *}
+ \<open>\<guillemotleft>\<dots>\<guillemotright>\<dots>\<close> to indicate where abstractions occur.\<close>
nominal_datatype ty =
Tvar "tyvrs"
@@ -46,10 +46,10 @@
| App "trm" "trm" (infixl "\<cdot>" 200)
| TApp "trm" "ty" (infixl "\<cdot>\<^sub>\<tau>" 200)
-text {* To be polite to the eye, some more familiar notation is introduced.
+text \<open>To be polite to the eye, some more familiar notation is introduced.
Because of the change in the order of arguments, one needs to use
translation rules, instead of syntax annotations at the term-constructors
- as given above for @{term "Arrow"}. *}
+ as given above for @{term "Arrow"}.\<close>
abbreviation
Forall_syn :: "tyvrs \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> ty" ("(3\<forall>_<:_./ _)" [0, 0, 10] 10)
@@ -66,19 +66,19 @@
where
"\<lambda>X<:T. t \<equiv> trm.TAbs X t T"
-text {* Again there are numerous facts that are proved automatically for @{typ "ty"}
- and @{typ "trm"}: for example that the set of free variables, i.e.~the @{text "support"},
+text \<open>Again there are numerous facts that are proved automatically for @{typ "ty"}
+ and @{typ "trm"}: for example that the set of free variables, i.e.~the \<open>support\<close>,
is finite. However note that nominal-datatype declarations do \emph{not} define
``classical" constructor-based datatypes, but rather define $\alpha$-equivalence
classes---we can for example show that $\alpha$-equivalent @{typ "ty"}s
- and @{typ "trm"}s are equal: *}
+ and @{typ "trm"}s are equal:\<close>
lemma alpha_illustration:
shows "(\<forall>X<:T. Tvar X) = (\<forall>Y<:T. Tvar Y)"
and "(\<lambda>x:T. Var x) = (\<lambda>y:T. Var y)"
by (simp_all add: ty.inject trm.inject alpha calc_atm fresh_atm)
-section {* SubTyping Contexts *}
+section \<open>SubTyping Contexts\<close>
nominal_datatype binding =
VarB vrs ty
@@ -86,12 +86,12 @@
type_synonym env = "binding list"
-text {* Typing contexts are represented as lists that ``grow" on the left; we
+text \<open>Typing contexts are represented as lists that ``grow" on the left; we
thereby deviating from the convention in the POPLmark-paper. The lists contain
- pairs of type-variables and types (this is sufficient for Part 1A). *}
+ pairs of type-variables and types (this is sufficient for Part 1A).\<close>
-text {* In order to state validity-conditions for typing-contexts, the notion of
- a @{text "dom"} of a typing-context is handy. *}
+text \<open>In order to state validity-conditions for typing-contexts, the notion of
+ a \<open>dom\<close> of a typing-context is handy.\<close>
nominal_primrec
"tyvrs_of" :: "binding \<Rightarrow> tyvrs set"
@@ -218,11 +218,11 @@
apply(auto simp add: fresh_prod fresh_list_cons tyvrs_fresh)
done
-text {* Not all lists of type @{typ "env"} are well-formed. One condition
+text \<open>Not all lists of type @{typ "env"} are well-formed. One condition
requires that in @{term "TVarB X S#\<Gamma>"} all free variables of @{term "S"} must be
- in the @{term "ty_dom"} of @{term "\<Gamma>"}, that is @{term "S"} must be @{text "closed"}
+ in the @{term "ty_dom"} of @{term "\<Gamma>"}, that is @{term "S"} must be \<open>closed\<close>
in @{term "\<Gamma>"}. The set of free variables of @{term "S"} is the
- @{text "support"} of @{term "S"}. *}
+ \<open>support\<close> of @{term "S"}.\<close>
definition "closed_in" :: "ty \<Rightarrow> env \<Rightarrow> bool" ("_ closed'_in _" [100,100] 100) where
"S closed_in \<Gamma> \<equiv> (supp S)\<subseteq>(ty_dom \<Gamma>)"
@@ -284,7 +284,7 @@
lemma closed_in_fresh: "(X::tyvrs) \<sharp> ty_dom \<Gamma> \<Longrightarrow> T closed_in \<Gamma> \<Longrightarrow> X \<sharp> T"
by (auto simp add: closed_in_def fresh_def ty_dom_supp)
-text {* Now validity of a context is a straightforward inductive definition. *}
+text \<open>Now validity of a context is a straightforward inductive definition.\<close>
inductive
valid_rel :: "env \<Rightarrow> bool" ("\<turnstile> _ ok" [100] 100)
@@ -332,7 +332,7 @@
(auto elim: validE intro!: valid_cons simp add: doms_append closed_in_def)
qed
-text {* Well-formed contexts have a unique type-binding for a type-variable. *}
+text \<open>Well-formed contexts have a unique type-binding for a type-variable.\<close>
lemma uniqueness_of_ctxt:
fixes \<Gamma>::"env"
@@ -383,7 +383,7 @@
ultimately show "T=S" by (auto simp add: binding.inject)
qed (auto)
-section {* Size and Capture-Avoiding Substitution for Types *}
+section \<open>Size and Capture-Avoiding Substitution for Types\<close>
nominal_primrec
size_ty :: "ty \<Rightarrow> nat"
@@ -590,14 +590,14 @@
by (nominal_induct t avoiding: X Y U rule: trm.strong_induct)
(simp_all add: fresh_atm calc_atm abs_fresh fresh_aux type_subst_rename)
-section {* Subtyping-Relation *}
+section \<open>Subtyping-Relation\<close>
-text {* The definition for the subtyping-relation follows quite closely what is written
+text \<open>The definition for the subtyping-relation follows quite closely what is written
in the POPLmark-paper, except for the premises dealing with well-formed contexts and
- the freshness constraint @{term "X\<sharp>\<Gamma>"} in the @{text "S_Forall"}-rule. (The freshness
+ the freshness constraint @{term "X\<sharp>\<Gamma>"} in the \<open>S_Forall\<close>-rule. (The freshness
constraint is specific to the \emph{nominal approach}. Note, however, that the constraint
does \emph{not} make the subtyping-relation ``partial"\ldots because we work over
- $\alpha$-equivalence classes.) *}
+ $\alpha$-equivalence classes.)\<close>
inductive
subtype_of :: "env \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> bool" ("_\<turnstile>_<:_" [100,100,100] 100)
@@ -670,7 +670,7 @@
apply (force simp add: closed_in_fresh dest: subtype_implies_closed subtype_implies_ok)+
done
-section {* Reflexivity of Subtyping *}
+section \<open>Reflexivity of Subtyping\<close>
lemma subtype_reflexivity:
assumes a: "\<turnstile> \<Gamma> ok"
@@ -702,18 +702,18 @@
using a b
apply(nominal_induct T avoiding: \<Gamma> rule: ty.strong_induct)
apply(auto simp add: ty.supp abs_supp supp_atm closed_in_def)
- --{* Too bad that this instantiation cannot be found automatically by
+ \<comment>\<open>Too bad that this instantiation cannot be found automatically by
\isakeyword{auto}; \isakeyword{blast} would find it if we had not used
- an explicit definition for @{text "closed_in_def"}. *}
+ an explicit definition for \<open>closed_in_def\<close>.\<close>
apply(drule_tac x="(TVarB tyvrs ty2)#\<Gamma>" in meta_spec)
apply(force dest: fresh_dom simp add: closed_in_def)
done
-section {* Weakening *}
+section \<open>Weakening\<close>
-text {* In order to prove weakening we introduce the notion of a type-context extending
+text \<open>In order to prove weakening we introduce the notion of a type-context extending
another. This generalization seems to make the proof for weakening to be
- smoother than if we had strictly adhered to the version in the POPLmark-paper. *}
+ smoother than if we had strictly adhered to the version in the POPLmark-paper.\<close>
definition extends :: "env \<Rightarrow> env \<Rightarrow> bool" ("_ extends _" [100,100] 100) where
"\<Delta> extends \<Gamma> \<equiv> \<forall>X Q. (TVarB X Q)\<in>set \<Gamma> \<longrightarrow> (TVarB X Q)\<in>set \<Delta>"
@@ -794,7 +794,7 @@
ultimately show "\<Delta> \<turnstile> (\<forall>X<:S\<^sub>1. S\<^sub>2) <: (\<forall>X<:T\<^sub>1. T\<^sub>2)" using ok by (force intro: SA_all)
qed
-text {* In fact all ``non-binding" cases can be solved automatically: *}
+text \<open>In fact all ``non-binding" cases can be solved automatically:\<close>
lemma weakening_more_automated:
assumes a: "\<Gamma> \<turnstile> S <: T"
@@ -822,9 +822,9 @@
ultimately show "\<Delta> \<turnstile> (\<forall>X<:S\<^sub>1. S\<^sub>2) <: (\<forall>X<:T\<^sub>1. T\<^sub>2)" using ok by (force intro: SA_all)
qed (blast intro: extends_closed extends_memb dest: extends_ty_dom)+
-section {* Transitivity and Narrowing *}
+section \<open>Transitivity and Narrowing\<close>
-text {* Some inversion lemmas that are needed in the transitivity and narrowing proof.*}
+text \<open>Some inversion lemmas that are needed in the transitivity and narrowing proof.\<close>
declare ty.inject [simp add]
@@ -840,14 +840,14 @@
apply(auto simp add: abs_fresh ty.inject alpha)
done
-text {* Next we prove the transitivity and narrowing for the subtyping-relation.
+text \<open>Next we prove the transitivity and narrowing for the subtyping-relation.
The POPLmark-paper says the following:
\begin{quote}
\begin{lemma}[Transitivity and Narrowing] \
\begin{enumerate}
\item If @{term "\<Gamma> \<turnstile> S<:Q"} and @{term "\<Gamma> \<turnstile> Q<:T"}, then @{term "\<Gamma> \<turnstile> S<:T"}.
-\item If @{text "\<Gamma>,X<:Q,\<Delta> \<turnstile> M<:N"} and @{term "\<Gamma> \<turnstile> P<:Q"} then @{text "\<Gamma>,X<:P,\<Delta> \<turnstile> M<:N"}.
+\item If \<open>\<Gamma>,X<:Q,\<Delta> \<turnstile> M<:N\<close> and @{term "\<Gamma> \<turnstile> P<:Q"} then \<open>\<Gamma>,X<:P,\<Delta> \<turnstile> M<:N\<close>.
\end{enumerate}
\end{lemma}
@@ -858,7 +858,7 @@
\end{quote}
For the induction on the size of @{term "Q"}, we use the induction-rule
-@{text "measure_induct_rule"}:
+\<open>measure_induct_rule\<close>:
\begin{center}
@{thm measure_induct_rule[of "size_ty",no_vars]}
@@ -867,7 +867,7 @@
That means in order to show a property @{term "P a"} for all @{term "a"},
the induct-rule requires to prove that for all @{term x} @{term "P x"} holds using the
assumption that for all @{term y} whose size is strictly smaller than
-that of @{term x} the property @{term "P y"} holds. *}
+that of @{term x} the property @{term "P y"} holds.\<close>
lemma
shows subtype_transitivity: "\<Gamma>\<turnstile>S<:Q \<Longrightarrow> \<Gamma>\<turnstile>Q<:T \<Longrightarrow> \<Gamma>\<turnstile>S<:T"
@@ -886,7 +886,7 @@
case (SA_Top \<Gamma> S)
then have rh_drv: "\<Gamma> \<turnstile> Top <: T" by simp
then have T_inst: "T = Top" by (auto elim: S_TopE)
- from `\<turnstile> \<Gamma> ok` and `S closed_in \<Gamma>`
+ from \<open>\<turnstile> \<Gamma> ok\<close> and \<open>S closed_in \<Gamma>\<close>
have "\<Gamma> \<turnstile> S <: Top" by auto
then show "\<Gamma> \<turnstile> S <: T" using T_inst by simp
next
@@ -900,7 +900,7 @@
next
case (SA_arrow \<Gamma> Q\<^sub>1 S\<^sub>1 S\<^sub>2 Q\<^sub>2)
then have rh_drv: "\<Gamma> \<turnstile> Q\<^sub>1 \<rightarrow> Q\<^sub>2 <: T" by simp
- from `Q\<^sub>1 \<rightarrow> Q\<^sub>2 = Q`
+ from \<open>Q\<^sub>1 \<rightarrow> Q\<^sub>2 = Q\<close>
have Q\<^sub>12_less: "size_ty Q\<^sub>1 < size_ty Q" "size_ty Q\<^sub>2 < size_ty Q" by auto
have lh_drv_prm\<^sub>1: "\<Gamma> \<turnstile> Q\<^sub>1 <: S\<^sub>1" by fact
have lh_drv_prm\<^sub>2: "\<Gamma> \<turnstile> S\<^sub>2 <: Q\<^sub>2" by fact
@@ -969,23 +969,23 @@
qed
} note transitivity_lemma = this
- { --{* The transitivity proof is now by the auxiliary lemma. *}
+ { \<comment>\<open>The transitivity proof is now by the auxiliary lemma.\<close>
case 1
- from `\<Gamma> \<turnstile> S <: Q` and `\<Gamma> \<turnstile> Q <: T`
+ from \<open>\<Gamma> \<turnstile> S <: Q\<close> and \<open>\<Gamma> \<turnstile> Q <: T\<close>
show "\<Gamma> \<turnstile> S <: T" by (rule transitivity_lemma)
next
case 2
- from `(\<Delta>@[(TVarB X Q)]@\<Gamma>) \<turnstile> M <: N`
- and `\<Gamma> \<turnstile> P<:Q`
+ from \<open>(\<Delta>@[(TVarB X Q)]@\<Gamma>) \<turnstile> M <: N\<close>
+ and \<open>\<Gamma> \<turnstile> P<:Q\<close>
show "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> M <: N"
proof (induct "\<Delta>@[(TVarB X Q)]@\<Gamma>" M N arbitrary: \<Gamma> X \<Delta> rule: subtype_of.induct)
case (SA_Top S \<Gamma> X \<Delta>)
- from `\<Gamma> \<turnstile> P <: Q`
+ from \<open>\<Gamma> \<turnstile> P <: Q\<close>
have "P closed_in \<Gamma>" by (simp add: subtype_implies_closed)
- with `\<turnstile> (\<Delta>@[(TVarB X Q)]@\<Gamma>) ok` have "\<turnstile> (\<Delta>@[(TVarB X P)]@\<Gamma>) ok"
+ with \<open>\<turnstile> (\<Delta>@[(TVarB X Q)]@\<Gamma>) ok\<close> have "\<turnstile> (\<Delta>@[(TVarB X P)]@\<Gamma>) ok"
by (simp add: replace_type)
moreover
- from `S closed_in (\<Delta>@[(TVarB X Q)]@\<Gamma>)` have "S closed_in (\<Delta>@[(TVarB X P)]@\<Gamma>)"
+ from \<open>S closed_in (\<Delta>@[(TVarB X Q)]@\<Gamma>)\<close> have "S closed_in (\<Delta>@[(TVarB X P)]@\<Gamma>)"
by (simp add: closed_in_def doms_append)
ultimately show "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> S <: Top" by (simp add: subtype_of.SA_Top)
next
@@ -1016,12 +1016,12 @@
qed
next
case (SA_refl_TVar Y \<Gamma> X \<Delta>)
- from `\<Gamma> \<turnstile> P <: Q`
+ from \<open>\<Gamma> \<turnstile> P <: Q\<close>
have "P closed_in \<Gamma>" by (simp add: subtype_implies_closed)
- with `\<turnstile> (\<Delta>@[(TVarB X Q)]@\<Gamma>) ok` have "\<turnstile> (\<Delta>@[(TVarB X P)]@\<Gamma>) ok"
+ with \<open>\<turnstile> (\<Delta>@[(TVarB X Q)]@\<Gamma>) ok\<close> have "\<turnstile> (\<Delta>@[(TVarB X P)]@\<Gamma>) ok"
by (simp add: replace_type)
moreover
- from `Y \<in> ty_dom (\<Delta>@[(TVarB X Q)]@\<Gamma>)` have "Y \<in> ty_dom (\<Delta>@[(TVarB X P)]@\<Gamma>)"
+ from \<open>Y \<in> ty_dom (\<Delta>@[(TVarB X Q)]@\<Gamma>)\<close> have "Y \<in> ty_dom (\<Delta>@[(TVarB X P)]@\<Gamma>)"
by (simp add: doms_append)
ultimately show "(\<Delta>@[(TVarB X P)]@\<Gamma>) \<turnstile> Tvar Y <: Tvar Y" by (simp add: subtype_of.SA_refl_TVar)
next
@@ -1037,7 +1037,7 @@
}
qed
-section {* Typing *}
+section \<open>Typing\<close>
inductive
typing :: "env \<Rightarrow> trm \<Rightarrow> ty \<Rightarrow> bool" ("_ \<turnstile> _ : _" [60,60,60] 60)
@@ -1116,22 +1116,22 @@
using assms
proof induct
case (T_Var x T \<Gamma>)
- from `\<turnstile> \<Gamma> ok` and `VarB x T \<in> set \<Gamma>`
+ from \<open>\<turnstile> \<Gamma> ok\<close> and \<open>VarB x T \<in> set \<Gamma>\<close>
show ?case by (rule ok_imp_VarB_closed_in)
next
case (T_App \<Gamma> t\<^sub>1 T\<^sub>1 T\<^sub>2 t\<^sub>2)
then show ?case by (auto simp add: ty.supp closed_in_def)
next
case (T_Abs x T\<^sub>1 \<Gamma> t\<^sub>2 T\<^sub>2)
- from `VarB x T\<^sub>1 # \<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>2`
+ from \<open>VarB x T\<^sub>1 # \<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>2\<close>
have "T\<^sub>1 closed_in \<Gamma>" by (auto dest: typing_ok)
with T_Abs show ?case by (auto simp add: ty.supp closed_in_def)
next
case (T_Sub \<Gamma> t S T)
- from `\<Gamma> \<turnstile> S <: T` show ?case by (simp add: subtype_implies_closed)
+ from \<open>\<Gamma> \<turnstile> S <: T\<close> show ?case by (simp add: subtype_implies_closed)
next
case (T_TAbs X T\<^sub>1 \<Gamma> t\<^sub>2 T\<^sub>2)
- from `TVarB X T\<^sub>1 # \<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>2`
+ from \<open>TVarB X T\<^sub>1 # \<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>2\<close>
have "T\<^sub>1 closed_in \<Gamma>" by (auto dest: typing_ok)
with T_TAbs show ?case by (auto simp add: ty.supp closed_in_def abs_supp)
next
@@ -1144,7 +1144,7 @@
qed
-subsection {* Evaluation *}
+subsection \<open>Evaluation\<close>
inductive
val :: "trm \<Rightarrow> bool"
@@ -1302,7 +1302,7 @@
qed
qed
-text {* A.5(6) *}
+text \<open>A.5(6)\<close>
lemma type_weaken:
assumes "(\<Delta>@\<Gamma>) \<turnstile> t : T"
@@ -1337,15 +1337,15 @@
then show ?case by (rule typing.T_Abs)
next
case (T_Sub t S T \<Delta> \<Gamma>)
- from refl and `\<turnstile> (\<Delta> @ B # \<Gamma>) ok`
+ from refl and \<open>\<turnstile> (\<Delta> @ B # \<Gamma>) ok\<close>
have "\<Delta> @ B # \<Gamma> \<turnstile> t : S" by (rule T_Sub)
- moreover from `(\<Delta> @ \<Gamma>)\<turnstile>S<:T` and `\<turnstile> (\<Delta> @ B # \<Gamma>) ok`
+ moreover from \<open>(\<Delta> @ \<Gamma>)\<turnstile>S<:T\<close> and \<open>\<turnstile> (\<Delta> @ B # \<Gamma>) ok\<close>
have "(\<Delta> @ B # \<Gamma>)\<turnstile>S<:T"
by (rule weakening) (simp add: extends_def T_Sub)
ultimately show ?case by (rule typing.T_Sub)
next
case (T_TAbs X T\<^sub>1 t\<^sub>2 T\<^sub>2 \<Delta> \<Gamma>)
- from `TVarB X T\<^sub>1 # \<Delta> @ \<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>2`
+ from \<open>TVarB X T\<^sub>1 # \<Delta> @ \<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>2\<close>
have closed: "T\<^sub>1 closed_in (\<Delta> @ \<Gamma>)"
by (auto dest: typing_ok)
have "\<turnstile> (TVarB X T\<^sub>1 # \<Delta> @ B # \<Gamma>) ok"
@@ -1367,7 +1367,7 @@
case (T_TApp X t\<^sub>1 T2 T11 T12 \<Delta> \<Gamma>)
have "\<Delta> @ B # \<Gamma> \<turnstile> t\<^sub>1 : (\<forall>X<:T11. T12)"
by (rule T_TApp refl)+
- moreover from `(\<Delta> @ \<Gamma>)\<turnstile>T2<:T11` and `\<turnstile> (\<Delta> @ B # \<Gamma>) ok`
+ moreover from \<open>(\<Delta> @ \<Gamma>)\<turnstile>T2<:T11\<close> and \<open>\<turnstile> (\<Delta> @ B # \<Gamma>) ok\<close>
have "(\<Delta> @ B # \<Gamma>)\<turnstile>T2<:T11"
by (rule weakening) (simp add: extends_def T_TApp)
ultimately show ?case by (rule better_T_TApp)
@@ -1382,7 +1382,7 @@
apply simp_all
done
-text {* A.6 *}
+text \<open>A.6\<close>
lemma strengthening:
assumes "(\<Gamma> @ VarB x Q # \<Delta>) \<turnstile> S <: T"
@@ -1395,7 +1395,7 @@
ultimately show ?case using subtype_of.SA_Top by auto
next
case (SA_refl_TVar X)
- from `\<turnstile> (\<Gamma> @ VarB x Q # \<Delta>) ok`
+ from \<open>\<turnstile> (\<Gamma> @ VarB x Q # \<Delta>) ok\<close>
have h1:"\<turnstile> (\<Gamma> @ \<Delta>) ok" by (auto dest: valid_cons')
have "X \<in> ty_dom (\<Gamma> @ VarB x Q # \<Delta>)" using SA_refl_TVar by auto
then have h2:"X \<in> ty_dom (\<Gamma> @ \<Delta>)" using ty_dom_vrs by auto
@@ -1407,7 +1407,7 @@
then show ?case using h1 h2 by auto
qed (auto)
-lemma narrow_type: -- {* A.7 *}
+lemma narrow_type: \<comment> \<open>A.7\<close>
assumes H: "\<Delta> @ (TVarB X Q) # \<Gamma> \<turnstile> t : T"
shows "\<Gamma> \<turnstile> P <: Q \<Longrightarrow> \<Delta> @ (TVarB X P) # \<Gamma> \<turnstile> t : T"
using H
@@ -1433,18 +1433,18 @@
case (T_TApp X' t1 T2 T11 T12 P D)
then have "D @ TVarB X P # \<Gamma> \<turnstile> t1 : Forall X' T12 T11" by fastforce
moreover have "(D @ [TVarB X Q] @ \<Gamma>) \<turnstile> T2<:T11" using T_TApp by auto
- then have "(D @ [TVarB X P] @ \<Gamma>) \<turnstile> T2<:T11" using `\<Gamma>\<turnstile>P<:Q`
+ then have "(D @ [TVarB X P] @ \<Gamma>) \<turnstile> T2<:T11" using \<open>\<Gamma>\<turnstile>P<:Q\<close>
by (rule subtype_narrow)
moreover from T_TApp have "X' \<sharp> (D @ TVarB X P # \<Gamma>, t1, T2)"
by (simp add: fresh_list_append fresh_list_cons fresh_prod)
ultimately show ?case by auto
qed
-subsection {* Substitution lemmas *}
+subsection \<open>Substitution lemmas\<close>
-subsubsection {* Substition Preserves Typing *}
+subsubsection \<open>Substition Preserves Typing\<close>
-theorem subst_type: -- {* A.8 *}
+theorem subst_type: \<comment> \<open>A.8\<close>
assumes H: "(\<Delta> @ (VarB x U) # \<Gamma>) \<turnstile> t : T"
shows "\<Gamma> \<turnstile> u : U \<Longrightarrow> \<Delta> @ \<Gamma> \<turnstile> t[x \<mapsto> u] : T" using H
proof (nominal_induct "\<Delta> @ (VarB x U) # \<Gamma>" t T avoiding: x u arbitrary: \<Delta> rule: typing.strong_induct)
@@ -1473,9 +1473,9 @@
ultimately show ?case by auto
next
case (T_TAbs X T1 t2 T2 x u D)
- from `TVarB X T1 # D @ VarB x U # \<Gamma> \<turnstile> t2 : T2` have "X \<sharp> T1"
+ from \<open>TVarB X T1 # D @ VarB x U # \<Gamma> \<turnstile> t2 : T2\<close> have "X \<sharp> T1"
by (auto simp add: valid_ty_dom_fresh dest: typing_ok intro!: closed_in_fresh)
- with `X \<sharp> u` and T_TAbs show ?case by fastforce
+ with \<open>X \<sharp> u\<close> and T_TAbs show ?case by fastforce
next
case (T_TApp X t1 T2 T11 T12 x u D)
then have "(D@\<Gamma>) \<turnstile>T2<:T11" using T_TApp by (auto dest: strengthening)
@@ -1483,9 +1483,9 @@
by (force simp add: fresh_prod fresh_list_append fresh_list_cons subst_trm_fresh_tyvar)
qed
-subsubsection {* Type Substitution Preserves Subtyping *}
+subsubsection \<open>Type Substitution Preserves Subtyping\<close>
-lemma substT_subtype: -- {* A.10 *}
+lemma substT_subtype: \<comment> \<open>A.10\<close>
assumes H: "(\<Delta> @ ((TVarB X Q) # \<Gamma>)) \<turnstile> S <: T"
shows "\<Gamma> \<turnstile> P <: Q \<Longrightarrow> (\<Delta>[X \<mapsto> P]\<^sub>e @ \<Gamma>) \<turnstile> S[X \<mapsto> P]\<^sub>\<tau> <: T[X \<mapsto> P]\<^sub>\<tau>"
using H
@@ -1508,11 +1508,11 @@
proof (cases "X = Y")
assume eq: "X = Y"
from eq and SA_trans_TVar have "TVarB Y Q \<in> set (D @ TVarB X Q # \<Gamma>)" by simp
- with G_ok have QS: "Q = S" using `TVarB Y S \<in> set (D @ TVarB X Q # \<Gamma>)`
+ with G_ok have QS: "Q = S" using \<open>TVarB Y S \<in> set (D @ TVarB X Q # \<Gamma>)\<close>
by (rule uniqueness_of_ctxt)
from X_\<Gamma>_ok have "X \<sharp> ty_dom \<Gamma>" and "Q closed_in \<Gamma>" by auto
then have XQ: "X \<sharp> Q" by (rule closed_in_fresh)
- note `\<Gamma>\<turnstile>P<:Q`
+ note \<open>\<Gamma>\<turnstile>P<:Q\<close>
moreover from ST have "\<turnstile> (D[X \<mapsto> P]\<^sub>e @ \<Gamma>) ok" by (rule subtype_implies_ok)
moreover have "(D[X \<mapsto> P]\<^sub>e @ \<Gamma>) extends \<Gamma>" by (simp add: extends_def)
ultimately have "(D[X \<mapsto> P]\<^sub>e @ \<Gamma>) \<turnstile> P<:Q" by (rule weakening)
@@ -1547,7 +1547,7 @@
qed
next
case (SA_refl_TVar Y X P D)
- note `\<turnstile> (D @ TVarB X Q # \<Gamma>) ok`
+ note \<open>\<turnstile> (D @ TVarB X Q # \<Gamma>) ok\<close>
moreover from SA_refl_TVar have closed: "P closed_in \<Gamma>"
by (auto dest: subtype_implies_closed)
ultimately have ok: "\<turnstile> (D[X \<mapsto> P]\<^sub>e @ \<Gamma>) ok" using valid_subst by auto
@@ -1582,9 +1582,9 @@
with SA_all and S1 show ?case by force
qed
-subsubsection {* Type Substitution Preserves Typing *}
+subsubsection \<open>Type Substitution Preserves Typing\<close>
-theorem substT_type: -- {* A.11 *}
+theorem substT_type: \<comment> \<open>A.11\<close>
assumes H: "(D @ TVarB X Q # G) \<turnstile> t : T"
shows "G \<turnstile> P <: Q \<Longrightarrow>
(D[X \<mapsto> P]\<^sub>e @ G) \<turnstile> t[X \<mapsto>\<^sub>\<tau> P] : T[X \<mapsto> P]\<^sub>\<tau>" using H
@@ -1592,9 +1592,9 @@
case (T_Var x T X P D')
have "G\<turnstile>P<:Q" by fact
then have "P closed_in G" using subtype_implies_closed by auto
- moreover note `\<turnstile> (D' @ TVarB X Q # G) ok`
+ moreover note \<open>\<turnstile> (D' @ TVarB X Q # G) ok\<close>
ultimately have "\<turnstile> (D'[X \<mapsto> P]\<^sub>e @ G) ok" using valid_subst by auto
- moreover note `VarB x T \<in> set (D' @ TVarB X Q # G)`
+ moreover note \<open>VarB x T \<in> set (D' @ TVarB X Q # G)\<close>
then have "VarB x T \<in> set D' \<or> VarB x T \<in> set G" by simp
then have "(VarB x (T[X \<mapsto> P]\<^sub>\<tau>)) \<in> set (D'[X \<mapsto> P]\<^sub>e @ G)"
proof
@@ -1648,7 +1648,7 @@
intro: substT_subtype)
qed
-lemma Abs_type: -- {* A.13(1) *}
+lemma Abs_type: \<comment> \<open>A.13(1)\<close>
assumes H: "\<Gamma> \<turnstile> (\<lambda>x:S. s) : T"
and H': "\<Gamma> \<turnstile> T <: U \<rightarrow> U'"
and H'': "x \<sharp> \<Gamma>"
@@ -1658,7 +1658,7 @@
using H H' H''
proof (nominal_induct \<Gamma> t \<equiv> "\<lambda>x:S. s" T avoiding: x arbitrary: U U' S s rule: typing.strong_induct)
case (T_Abs y T\<^sub>1 \<Gamma> t\<^sub>2 T\<^sub>2)
- from `\<Gamma> \<turnstile> T\<^sub>1 \<rightarrow> T\<^sub>2 <: U \<rightarrow> U'`
+ from \<open>\<Gamma> \<turnstile> T\<^sub>1 \<rightarrow> T\<^sub>2 <: U \<rightarrow> U'\<close>
obtain ty1: "\<Gamma> \<turnstile> U <: S" and ty2: "\<Gamma> \<turnstile> T\<^sub>2 <: U'" using T_Abs
by cases (simp_all add: ty.inject trm.inject alpha fresh_atm)
from T_Abs have "VarB y S # \<Gamma> \<turnstile> [(y, x)] \<bullet> s : T\<^sub>2"
@@ -1690,7 +1690,7 @@
using H subtype_reflexivity_from_typing [OF H] H'
by (rule Abs_type)
-lemma TAbs_type: -- {* A.13(2) *}
+lemma TAbs_type: \<comment> \<open>A.13(2)\<close>
assumes H: "\<Gamma> \<turnstile> (\<lambda>X<:S. s) : T"
and H': "\<Gamma> \<turnstile> T <: (\<forall>X<:U. U')"
and fresh: "X \<sharp> \<Gamma>" "X \<sharp> S" "X \<sharp> U"
@@ -1701,9 +1701,9 @@
using H H' fresh
proof (nominal_induct \<Gamma> t \<equiv> "\<lambda>X<:S. s" T avoiding: X U U' S arbitrary: s rule: typing.strong_induct)
case (T_TAbs Y T\<^sub>1 \<Gamma> t\<^sub>2 T\<^sub>2)
- from `TVarB Y T\<^sub>1 # \<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>2` have Y: "Y \<sharp> \<Gamma>"
+ from \<open>TVarB Y T\<^sub>1 # \<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>2\<close> have Y: "Y \<sharp> \<Gamma>"
by (auto dest!: typing_ok simp add: valid_ty_dom_fresh)
- from `Y \<sharp> U'` and `Y \<sharp> X`
+ from \<open>Y \<sharp> U'\<close> and \<open>Y \<sharp> X\<close>
have "(\<forall>X<:U. U') = (\<forall>Y<:U. [(Y, X)] \<bullet> U')"
by (simp add: ty.inject alpha' fresh_atm)
with T_TAbs have "\<Gamma> \<turnstile> (\<forall>Y<:S. T\<^sub>2) <: (\<forall>Y<:U. [(Y, X)] \<bullet> U')" by (simp add: trm.inject)
@@ -1714,13 +1714,13 @@
by (simp add: trm.inject alpha fresh_atm)
then have "[(Y, X)] \<bullet> (TVarB Y S # \<Gamma>) \<turnstile> [(Y, X)] \<bullet> [(Y, X)] \<bullet> s : [(Y, X)] \<bullet> T\<^sub>2"
by (rule typing.eqvt)
- with `X \<sharp> \<Gamma>` `X \<sharp> S` Y `Y \<sharp> S` have "TVarB X S # \<Gamma> \<turnstile> s : [(Y, X)] \<bullet> T\<^sub>2"
+ with \<open>X \<sharp> \<Gamma>\<close> \<open>X \<sharp> S\<close> Y \<open>Y \<sharp> S\<close> have "TVarB X S # \<Gamma> \<turnstile> s : [(Y, X)] \<bullet> T\<^sub>2"
by perm_simp
then have "TVarB X U # \<Gamma> \<turnstile> s : [(Y, X)] \<bullet> T\<^sub>2" using ty1
by (rule narrow_type [of "[]", simplified])
moreover from ty2 have "([(Y, X)] \<bullet> (TVarB Y U # \<Gamma>)) \<turnstile> ([(Y, X)] \<bullet> T\<^sub>2) <: ([(Y, X)] \<bullet> [(Y, X)] \<bullet> U')"
by (rule subtype_of.eqvt)
- with `X \<sharp> \<Gamma>` `X \<sharp> U` Y `Y \<sharp> U` have "(TVarB X U # \<Gamma>) \<turnstile> ([(Y, X)] \<bullet> T\<^sub>2) <: U'"
+ with \<open>X \<sharp> \<Gamma>\<close> \<open>X \<sharp> U\<close> Y \<open>Y \<sharp> U\<close> have "(TVarB X U # \<Gamma>) \<turnstile> ([(Y, X)] \<bullet> T\<^sub>2) <: U'"
by perm_simp
ultimately show ?case by (rule T_TAbs)
next
@@ -1738,7 +1738,7 @@
using H subtype_reflexivity_from_typing [OF H] fresh
by (rule TAbs_type)
-theorem preservation: -- {* A.20 *}
+theorem preservation: \<comment> \<open>A.20\<close>
assumes H: "\<Gamma> \<turnstile> t : T"
shows "t \<longmapsto> t' \<Longrightarrow> \<Gamma> \<turnstile> t' : T" using H
proof (nominal_induct avoiding: t' rule: typing.strong_induct)
@@ -1747,7 +1747,7 @@
by (rule exists_fresh) (rule fin_supp)
obtain X::tyvrs where "X \<sharp> (t\<^sub>1 \<cdot> t\<^sub>2, t')"
by (rule exists_fresh) (rule fin_supp)
- with `t\<^sub>1 \<cdot> t\<^sub>2 \<longmapsto> t'` show ?case
+ with \<open>t\<^sub>1 \<cdot> t\<^sub>2 \<longmapsto> t'\<close> show ?case
proof (cases rule: eval.strong_cases [where x=x and X=X])
case (E_Abs v\<^sub>2 T\<^sub>11' t\<^sub>12)
with T_App and x_fresh have h: "\<Gamma> \<turnstile> (\<lambda>x:T\<^sub>11'. t\<^sub>12) : T\<^sub>11 \<rightarrow> T\<^sub>12"
@@ -1758,7 +1758,7 @@
and t\<^sub>12: "(VarB x T\<^sub>11') # \<Gamma> \<turnstile> t\<^sub>12 : S'"
and S': "\<Gamma> \<turnstile> S' <: T\<^sub>12"
by (rule Abs_type') blast
- from `\<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>11`
+ from \<open>\<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>11\<close>
have "\<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>11'" using T\<^sub>11 by (rule T_Sub)
with t\<^sub>12 have "\<Gamma> \<turnstile> t\<^sub>12[x \<mapsto> t\<^sub>2] : S'"
by (rule subst_type [where \<Delta>="[]", simplified])
@@ -1768,7 +1768,7 @@
case (E_App1 t''' t'' u)
hence "t\<^sub>1 \<longmapsto> t''" by (simp add:trm.inject)
hence "\<Gamma> \<turnstile> t'' : T\<^sub>11 \<rightarrow> T\<^sub>12" by (rule T_App)
- hence "\<Gamma> \<turnstile> t'' \<cdot> t\<^sub>2 : T\<^sub>12" using `\<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>11`
+ hence "\<Gamma> \<turnstile> t'' \<cdot> t\<^sub>2 : T\<^sub>12" using \<open>\<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>11\<close>
by (rule typing.T_App)
with E_App1 show ?thesis by (simp add:trm.inject)
next
@@ -1783,27 +1783,27 @@
case (T_TApp X \<Gamma> t\<^sub>1 T\<^sub>2 T\<^sub>11 T\<^sub>12 t')
obtain x::vrs where "x \<sharp> (t\<^sub>1 \<cdot>\<^sub>\<tau> T\<^sub>2, t')"
by (rule exists_fresh) (rule fin_supp)
- with `t\<^sub>1 \<cdot>\<^sub>\<tau> T\<^sub>2 \<longmapsto> t'`
+ with \<open>t\<^sub>1 \<cdot>\<^sub>\<tau> T\<^sub>2 \<longmapsto> t'\<close>
show ?case
proof (cases rule: eval.strong_cases [where X=X and x=x])
case (E_TAbs T\<^sub>11' T\<^sub>2' t\<^sub>12)
with T_TApp have "\<Gamma> \<turnstile> (\<lambda>X<:T\<^sub>11'. t\<^sub>12) : (\<forall>X<:T\<^sub>11. T\<^sub>12)" and "X \<sharp> \<Gamma>" and "X \<sharp> T\<^sub>11'"
by (simp_all add: trm.inject)
- moreover from `\<Gamma>\<turnstile>T\<^sub>2<:T\<^sub>11` and `X \<sharp> \<Gamma>` have "X \<sharp> T\<^sub>11"
+ moreover from \<open>\<Gamma>\<turnstile>T\<^sub>2<:T\<^sub>11\<close> and \<open>X \<sharp> \<Gamma>\<close> have "X \<sharp> T\<^sub>11"
by (blast intro: closed_in_fresh fresh_dom dest: subtype_implies_closed)
ultimately obtain S'
where "TVarB X T\<^sub>11 # \<Gamma> \<turnstile> t\<^sub>12 : S'"
and "(TVarB X T\<^sub>11 # \<Gamma>) \<turnstile> S' <: T\<^sub>12"
by (rule TAbs_type') blast
hence "TVarB X T\<^sub>11 # \<Gamma> \<turnstile> t\<^sub>12 : T\<^sub>12" by (rule T_Sub)
- hence "\<Gamma> \<turnstile> t\<^sub>12[X \<mapsto>\<^sub>\<tau> T\<^sub>2] : T\<^sub>12[X \<mapsto> T\<^sub>2]\<^sub>\<tau>" using `\<Gamma> \<turnstile> T\<^sub>2 <: T\<^sub>11`
+ hence "\<Gamma> \<turnstile> t\<^sub>12[X \<mapsto>\<^sub>\<tau> T\<^sub>2] : T\<^sub>12[X \<mapsto> T\<^sub>2]\<^sub>\<tau>" using \<open>\<Gamma> \<turnstile> T\<^sub>2 <: T\<^sub>11\<close>
by (rule substT_type [where D="[]", simplified])
with T_TApp and E_TAbs show ?thesis by (simp add: trm.inject)
next
case (E_TApp t''' t'' T)
from E_TApp have "t\<^sub>1 \<longmapsto> t''" by (simp add: trm.inject)
then have "\<Gamma> \<turnstile> t'' : (\<forall>X<:T\<^sub>11. T\<^sub>12)" by (rule T_TApp)
- then have "\<Gamma> \<turnstile> t'' \<cdot>\<^sub>\<tau> T\<^sub>2 : T\<^sub>12[X \<mapsto> T\<^sub>2]\<^sub>\<tau>" using `\<Gamma> \<turnstile> T\<^sub>2 <: T\<^sub>11`
+ then have "\<Gamma> \<turnstile> t'' \<cdot>\<^sub>\<tau> T\<^sub>2 : T\<^sub>12[X \<mapsto> T\<^sub>2]\<^sub>\<tau>" using \<open>\<Gamma> \<turnstile> T\<^sub>2 <: T\<^sub>11\<close>
by (rule better_T_TApp)
with E_TApp show ?thesis by (simp add: trm.inject)
qed (simp_all add: fresh_prod)
@@ -1815,24 +1815,24 @@
ultimately show ?case by (rule typing.T_Sub)
qed (auto)
-lemma Fun_canonical: -- {* A.14(1) *}
+lemma Fun_canonical: \<comment> \<open>A.14(1)\<close>
assumes ty: "[] \<turnstile> v : T\<^sub>1 \<rightarrow> T\<^sub>2"
shows "val v \<Longrightarrow> \<exists>x t S. v = (\<lambda>x:S. t)" using ty
proof (induct "[]::env" v "T\<^sub>1 \<rightarrow> T\<^sub>2" arbitrary: T\<^sub>1 T\<^sub>2)
case (T_Sub t S)
- from `[] \<turnstile> S <: T\<^sub>1 \<rightarrow> T\<^sub>2`
+ from \<open>[] \<turnstile> S <: T\<^sub>1 \<rightarrow> T\<^sub>2\<close>
obtain S\<^sub>1 S\<^sub>2 where S: "S = S\<^sub>1 \<rightarrow> S\<^sub>2"
by cases (auto simp add: T_Sub)
- then show ?case using `val t` by (rule T_Sub)
+ then show ?case using \<open>val t\<close> by (rule T_Sub)
qed (auto)
-lemma TyAll_canonical: -- {* A.14(3) *}
+lemma TyAll_canonical: \<comment> \<open>A.14(3)\<close>
fixes X::tyvrs
assumes ty: "[] \<turnstile> v : (\<forall>X<:T\<^sub>1. T\<^sub>2)"
shows "val v \<Longrightarrow> \<exists>X t S. v = (\<lambda>X<:S. t)" using ty
proof (induct "[]::env" v "\<forall>X<:T\<^sub>1. T\<^sub>2" arbitrary: X T\<^sub>1 T\<^sub>2)
case (T_Sub t S)
- from `[] \<turnstile> S <: (\<forall>X<:T\<^sub>1. T\<^sub>2)`
+ from \<open>[] \<turnstile> S <: (\<forall>X<:T\<^sub>1. T\<^sub>2)\<close>
obtain X S\<^sub>1 S\<^sub>2 where S: "S = (\<forall>X<:S\<^sub>1. S\<^sub>2)"
by cases (auto simp add: T_Sub)
then show ?case using T_Sub by auto
--- a/src/HOL/Nominal/Examples/Height.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Nominal/Examples/Height.thy Thu May 26 17:51:22 2016 +0200
@@ -2,10 +2,10 @@
imports "../Nominal"
begin
-text {*
+text \<open>
A small problem suggested by D. Wang. It shows how
the height of a lambda-terms behaves under substitution.
-*}
+\<close>
atom_decl name
@@ -14,7 +14,7 @@
| App "lam" "lam"
| Lam "\<guillemotleft>name\<guillemotright>lam" ("Lam [_]._" [100,100] 100)
-text {* Definition of the height-function on lambda-terms. *}
+text \<open>Definition of the height-function on lambda-terms.\<close>
nominal_primrec
height :: "lam \<Rightarrow> int"
@@ -28,7 +28,7 @@
apply(fresh_guess add: perm_int_def)+
done
-text {* Definition of capture-avoiding substitution. *}
+text \<open>Definition of capture-avoiding substitution.\<close>
nominal_primrec
subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam" ("_[_::=_]" [100,100,100] 100)
@@ -42,16 +42,16 @@
apply(fresh_guess)+
done
-text{* The next lemma is needed in the Var-case of the theorem below. *}
+text\<open>The next lemma is needed in the Var-case of the theorem below.\<close>
lemma height_ge_one:
shows "1 \<le> (height e)"
by (nominal_induct e rule: lam.strong_induct) (simp_all)
-text {*
+text \<open>
Unlike the proplem suggested by Wang, however, the
theorem is here formulated entirely by using functions.
-*}
+\<close>
theorem height_subst:
shows "height (e[x::=e']) \<le> ((height e) - 1) + (height e')"
--- a/src/HOL/Nominal/Examples/Lam_Funs.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Nominal/Examples/Lam_Funs.thy Thu May 26 17:51:22 2016 +0200
@@ -2,10 +2,10 @@
imports "../Nominal"
begin
-text {*
+text \<open>
Provides useful definitions for reasoning
with lambda-terms.
-*}
+\<close>
atom_decl name
@@ -14,7 +14,7 @@
| App "lam" "lam"
| Lam "\<guillemotleft>name\<guillemotright>lam" ("Lam [_]._" [100,100] 100)
-text {* The depth of a lambda-term. *}
+text \<open>The depth of a lambda-term.\<close>
nominal_primrec
depth :: "lam \<Rightarrow> nat"
@@ -28,12 +28,12 @@
apply(fresh_guess)+
done
-text {*
+text \<open>
The free variables of a lambda-term. A complication in this
function arises from the fact that it returns a name set, which
is not a finitely supported type. Therefore we have to prove
the invariant that frees always returns a finite set of names.
-*}
+\<close>
nominal_primrec (invariant: "\<lambda>s::name set. finite s")
frees :: "lam \<Rightarrow> name set"
@@ -50,17 +50,17 @@
apply(fresh_guess)+
done
-text {*
+text \<open>
We can avoid the definition of frees by
using the build in notion of support.
-*}
+\<close>
lemma frees_equals_support:
shows "frees t = supp t"
by (nominal_induct t rule: lam.strong_induct)
(simp_all add: lam.supp supp_atm abs_supp)
-text {* Parallel and single capture-avoiding substitution. *}
+text \<open>Parallel and single capture-avoiding substitution.\<close>
fun
lookup :: "(name\<times>lam) list \<Rightarrow> name \<Rightarrow> lam"
@@ -112,18 +112,18 @@
apply(blast)+
done
-text {*
+text \<open>
Contexts - lambda-terms with a single hole.
Note that the lambda case in contexts does not bind a
name, even if we introduce the notation [_]._ for CLam.
-*}
+\<close>
nominal_datatype clam =
Hole ("\<box>" 1000)
| CAppL "clam" "lam"
| CAppR "lam" "clam"
| CLam "name" "clam" ("CLam [_]._" [100,100] 100)
-text {* Filling a lambda-term into a context. *}
+text \<open>Filling a lambda-term into a context.\<close>
nominal_primrec
filling :: "clam \<Rightarrow> lam \<Rightarrow> lam" ("_\<lbrakk>_\<rbrakk>" [100,100] 100)
@@ -134,7 +134,7 @@
| "(CLam [x].E)\<lbrakk>t\<rbrakk> = Lam [x].(E\<lbrakk>t\<rbrakk>)"
by (rule TrueI)+
-text {* Composition od two contexts *}
+text \<open>Composition od two contexts\<close>
nominal_primrec
clam_compose :: "clam \<Rightarrow> clam \<Rightarrow> clam" ("_ \<circ> _" [100,100] 100)
--- a/src/HOL/Nominal/Examples/Lambda_mu.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Nominal/Examples/Lambda_mu.thy Thu May 26 17:51:22 2016 +0200
@@ -2,7 +2,7 @@
imports "../Nominal"
begin
-section {* Lambda-Mu according to a paper by Gavin Bierman *}
+section \<open>Lambda-Mu according to a paper by Gavin Bierman\<close>
atom_decl var mvar
--- a/src/HOL/Nominal/Examples/LocalWeakening.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Nominal/Examples/LocalWeakening.thy Thu May 26 17:51:22 2016 +0200
@@ -9,17 +9,17 @@
atom_decl name
-text {*
+text \<open>
Curry-style lambda terms in locally nameless
representation without any binders
-*}
+\<close>
nominal_datatype llam =
lPar "name"
| lVar "nat"
| lApp "llam" "llam"
| lLam "llam"
-text {* definition of vsub - at the moment a bit cumbersome *}
+text \<open>definition of vsub - at the moment a bit cumbersome\<close>
lemma llam_cases:
"(\<exists>a. t = lPar a) \<or> (\<exists>n. t = lVar n) \<or> (\<exists>t1 t2. t = lApp t1 t2) \<or>
@@ -66,7 +66,7 @@
shows "pi\<bullet>(freshen t p) = freshen (pi\<bullet>t) (pi\<bullet>p)"
by (simp add: vsub_eqvt freshen_def perm_nat_def)
-text {* types *}
+text \<open>types\<close>
nominal_datatype ty =
TVar "nat"
@@ -79,7 +79,7 @@
by (induct T rule: ty.induct)
(simp_all add: fresh_nat)
-text {* valid contexts *}
+text \<open>valid contexts\<close>
type_synonym cxt = "(name\<times>ty) list"
@@ -96,7 +96,7 @@
shows "a\<sharp>\<Gamma> \<and> valid \<Gamma>"
using a by (cases) (auto)
-text {* "weak" typing relation *}
+text \<open>"weak" typing relation\<close>
inductive
typing :: "cxt\<Rightarrow>llam\<Rightarrow>ty\<Rightarrow>bool" (" _ \<turnstile> _ : _ " [60,60,60] 60)
@@ -112,18 +112,18 @@
shows "valid \<Gamma>"
using a by (induct) (auto dest: v2_elim)
-text {*
+text \<open>
we have to explicitly state that nominal_inductive should strengthen
over the variable x (since x is not a binder)
-*}
+\<close>
nominal_inductive typing
avoids t_lLam: x
by (auto simp add: fresh_prod dest: v2_elim typing_implies_valid)
-text {* strong induction principle for typing *}
+text \<open>strong induction principle for typing\<close>
thm typing.strong_induct
-text {* sub-contexts *}
+text \<open>sub-contexts\<close>
abbreviation
"sub_context" :: "cxt \<Rightarrow> cxt \<Rightarrow> bool" ("_ \<subseteq> _" [60,60] 60)
--- a/src/HOL/Nominal/Examples/Pattern.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Nominal/Examples/Pattern.thy Thu May 26 17:51:22 2016 +0200
@@ -1,4 +1,4 @@
-section {* Simply-typed lambda-calculus with let and tuple patterns *}
+section \<open>Simply-typed lambda-calculus with let and tuple patterns\<close>
theory Pattern
imports "../Nominal"
@@ -411,7 +411,7 @@
shows "\<Delta> @ \<Gamma> \<turnstile> t[x\<mapsto>u] : T" using a b
proof (nominal_induct \<Gamma>'\<equiv>"\<Delta> @ [(x, U)] @ \<Gamma>" t T avoiding: x u \<Delta> rule: typing.strong_induct)
case (Var y T x u \<Delta>)
- from `valid (\<Delta> @ [(x, U)] @ \<Gamma>)`
+ from \<open>valid (\<Delta> @ [(x, U)] @ \<Gamma>)\<close>
have valid: "valid (\<Delta> @ \<Gamma>)" by (rule valid_insert)
show "\<Delta> @ \<Gamma> \<turnstile> Var y[x\<mapsto>u] : T"
proof cases
@@ -425,19 +425,19 @@
qed
next
case (Tuple t\<^sub>1 T\<^sub>1 t\<^sub>2 T\<^sub>2)
- from refl `\<Gamma> \<turnstile> u : U`
+ from refl \<open>\<Gamma> \<turnstile> u : U\<close>
have "\<Delta> @ \<Gamma> \<turnstile> t\<^sub>1[x\<mapsto>u] : T\<^sub>1" by (rule Tuple)
- moreover from refl `\<Gamma> \<turnstile> u : U`
+ moreover from refl \<open>\<Gamma> \<turnstile> u : U\<close>
have "\<Delta> @ \<Gamma> \<turnstile> t\<^sub>2[x\<mapsto>u] : T\<^sub>2" by (rule Tuple)
ultimately have "\<Delta> @ \<Gamma> \<turnstile> \<langle>t\<^sub>1[x\<mapsto>u], t\<^sub>2[x\<mapsto>u]\<rangle> : T\<^sub>1 \<otimes> T\<^sub>2" ..
then show ?case by simp
next
case (Let p t T \<Delta>' s S)
- from refl `\<Gamma> \<turnstile> u : U`
+ from refl \<open>\<Gamma> \<turnstile> u : U\<close>
have "\<Delta> @ \<Gamma> \<turnstile> t[x\<mapsto>u] : T" by (rule Let)
- moreover note `\<turnstile> p : T \<Rightarrow> \<Delta>'`
+ moreover note \<open>\<turnstile> p : T \<Rightarrow> \<Delta>'\<close>
moreover have "\<Delta>' @ (\<Delta> @ [(x, U)] @ \<Gamma>) = (\<Delta>' @ \<Delta>) @ [(x, U)] @ \<Gamma>" by simp
- then have "(\<Delta>' @ \<Delta>) @ \<Gamma> \<turnstile> s[x\<mapsto>u] : S" using `\<Gamma> \<turnstile> u : U` by (rule Let)
+ then have "(\<Delta>' @ \<Delta>) @ \<Gamma> \<turnstile> s[x\<mapsto>u] : S" using \<open>\<Gamma> \<turnstile> u : U\<close> by (rule Let)
then have "\<Delta>' @ \<Delta> @ \<Gamma> \<turnstile> s[x\<mapsto>u] : S" by simp
ultimately have "\<Delta> @ \<Gamma> \<turnstile> (LET p = t[x\<mapsto>u] IN s[x\<mapsto>u]) : S"
by (rule better_T_Let)
@@ -448,7 +448,7 @@
case (Abs y T t S)
have "(y, T) # \<Delta> @ [(x, U)] @ \<Gamma> = ((y, T) # \<Delta>) @ [(x, U)] @ \<Gamma>"
by simp
- then have "((y, T) # \<Delta>) @ \<Gamma> \<turnstile> t[x\<mapsto>u] : S" using `\<Gamma> \<turnstile> u : U` by (rule Abs)
+ then have "((y, T) # \<Delta>) @ \<Gamma> \<turnstile> t[x\<mapsto>u] : S" using \<open>\<Gamma> \<turnstile> u : U\<close> by (rule Abs)
then have "(y, T) # \<Delta> @ \<Gamma> \<turnstile> t[x\<mapsto>u] : S" by simp
then have "\<Delta> @ \<Gamma> \<turnstile> (\<lambda>y:T. t[x\<mapsto>u]) : T \<rightarrow> S"
by (rule typing.Abs)
@@ -457,9 +457,9 @@
ultimately show ?case by simp
next
case (App t\<^sub>1 T S t\<^sub>2)
- from refl `\<Gamma> \<turnstile> u : U`
+ from refl \<open>\<Gamma> \<turnstile> u : U\<close>
have "\<Delta> @ \<Gamma> \<turnstile> t\<^sub>1[x\<mapsto>u] : T \<rightarrow> S" by (rule App)
- moreover from refl `\<Gamma> \<turnstile> u : U`
+ moreover from refl \<open>\<Gamma> \<turnstile> u : U\<close>
have "\<Delta> @ \<Gamma> \<turnstile> t\<^sub>2[x\<mapsto>u] : T" by (rule App)
ultimately have "\<Delta> @ \<Gamma> \<turnstile> (t\<^sub>1[x\<mapsto>u]) \<cdot> (t\<^sub>2[x\<mapsto>u]) : S"
by (rule typing.App)
@@ -489,14 +489,14 @@
shows "\<Gamma>\<^sub>1 @ \<Gamma>\<^sub>2 \<turnstile> \<theta>\<lparr>t\<rparr> : T" using assms
proof (induct arbitrary: \<Gamma>\<^sub>1 \<Gamma>\<^sub>2 t u T \<theta>)
case (PVar x U)
- from `\<Gamma>\<^sub>1 @ [(x, U)] @ \<Gamma>\<^sub>2 \<turnstile> t : T` `\<Gamma>\<^sub>2 \<turnstile> u : U`
+ from \<open>\<Gamma>\<^sub>1 @ [(x, U)] @ \<Gamma>\<^sub>2 \<turnstile> t : T\<close> \<open>\<Gamma>\<^sub>2 \<turnstile> u : U\<close>
have "\<Gamma>\<^sub>1 @ \<Gamma>\<^sub>2 \<turnstile> t[x\<mapsto>u] : T" by (rule subst_type_aux)
- moreover from `\<turnstile> PVar x U \<rhd> u \<Rightarrow> \<theta>` have "\<theta> = [(x, u)]"
+ moreover from \<open>\<turnstile> PVar x U \<rhd> u \<Rightarrow> \<theta>\<close> have "\<theta> = [(x, u)]"
by cases simp_all
ultimately show ?case by simp
next
case (PTuple p S \<Delta>\<^sub>1 q U \<Delta>\<^sub>2)
- from `\<turnstile> \<langle>\<langle>p, q\<rangle>\<rangle> \<rhd> u \<Rightarrow> \<theta>` obtain u\<^sub>1 u\<^sub>2 \<theta>\<^sub>1 \<theta>\<^sub>2
+ from \<open>\<turnstile> \<langle>\<langle>p, q\<rangle>\<rangle> \<rhd> u \<Rightarrow> \<theta>\<close> obtain u\<^sub>1 u\<^sub>2 \<theta>\<^sub>1 \<theta>\<^sub>2
where u: "u = \<langle>u\<^sub>1, u\<^sub>2\<rangle>" and \<theta>: "\<theta> = \<theta>\<^sub>1 @ \<theta>\<^sub>2"
and p: "\<turnstile> p \<rhd> u\<^sub>1 \<Rightarrow> \<theta>\<^sub>1" and q: "\<turnstile> q \<rhd> u\<^sub>2 \<Rightarrow> \<theta>\<^sub>2"
by cases simp_all
@@ -504,10 +504,10 @@
then obtain u\<^sub>1: "\<Gamma>\<^sub>2 \<turnstile> u\<^sub>1 : S" and u\<^sub>2: "\<Gamma>\<^sub>2 \<turnstile> u\<^sub>2 : U"
by cases (simp_all add: ty.inject trm.inject)
note u\<^sub>1
- moreover from `\<Gamma>\<^sub>1 @ (\<Delta>\<^sub>2 @ \<Delta>\<^sub>1) @ \<Gamma>\<^sub>2 \<turnstile> t : T`
+ moreover from \<open>\<Gamma>\<^sub>1 @ (\<Delta>\<^sub>2 @ \<Delta>\<^sub>1) @ \<Gamma>\<^sub>2 \<turnstile> t : T\<close>
have "(\<Gamma>\<^sub>1 @ \<Delta>\<^sub>2) @ \<Delta>\<^sub>1 @ \<Gamma>\<^sub>2 \<turnstile> t : T" by simp
moreover note p
- moreover from `supp \<langle>\<langle>p, q\<rangle>\<rangle> \<sharp>* u` and u
+ moreover from \<open>supp \<langle>\<langle>p, q\<rangle>\<rangle> \<sharp>* u\<close> and u
have "(supp p::name set) \<sharp>* u\<^sub>1" by (simp add: fresh_star_def)
ultimately have \<theta>\<^sub>1: "(\<Gamma>\<^sub>1 @ \<Delta>\<^sub>2) @ \<Gamma>\<^sub>2 \<turnstile> \<theta>\<^sub>1\<lparr>t\<rparr> : T"
by (rule PTuple)
@@ -515,11 +515,11 @@
moreover from \<theta>\<^sub>1
have "\<Gamma>\<^sub>1 @ \<Delta>\<^sub>2 @ \<Gamma>\<^sub>2 \<turnstile> \<theta>\<^sub>1\<lparr>t\<rparr> : T" by simp
moreover note q
- moreover from `supp \<langle>\<langle>p, q\<rangle>\<rangle> \<sharp>* u` and u
+ moreover from \<open>supp \<langle>\<langle>p, q\<rangle>\<rangle> \<sharp>* u\<close> and u
have "(supp q::name set) \<sharp>* u\<^sub>2" by (simp add: fresh_star_def)
ultimately have "\<Gamma>\<^sub>1 @ \<Gamma>\<^sub>2 \<turnstile> \<theta>\<^sub>2\<lparr>\<theta>\<^sub>1\<lparr>t\<rparr>\<rparr> : T"
by (rule PTuple)
- moreover from `\<turnstile> \<langle>\<langle>p, q\<rangle>\<rangle> \<rhd> u \<Rightarrow> \<theta>` `supp \<langle>\<langle>p, q\<rangle>\<rangle> \<sharp>* u`
+ moreover from \<open>\<turnstile> \<langle>\<langle>p, q\<rangle>\<rangle> \<rhd> u \<Rightarrow> \<theta>\<close> \<open>supp \<langle>\<langle>p, q\<rangle>\<rangle> \<sharp>* u\<close>
have "(supp (map fst \<theta>)::name set) \<sharp>* map snd \<theta>"
by (rule match_fresh)
ultimately show ?case using \<theta> by (simp add: psubst_append)
@@ -578,10 +578,10 @@
case (Abs x' T' t' U)
obtain y::name where y: "y \<sharp> (x, \<Gamma>, \<lambda>x':T'. t')"
by (rule exists_fresh) (auto intro: fin_supp)
- from `(\<lambda>x:T. t) = (\<lambda>x':T'. t')` [symmetric]
+ from \<open>(\<lambda>x:T. t) = (\<lambda>x':T'. t')\<close> [symmetric]
have x: "x \<sharp> (\<lambda>x':T'. t')" by (simp add: abs_fresh)
have x': "x' \<sharp> (\<lambda>x':T'. t')" by (simp add: abs_fresh)
- from `(x', T') # \<Gamma> \<turnstile> t' : U` have x'': "x' \<sharp> \<Gamma>"
+ from \<open>(x', T') # \<Gamma> \<turnstile> t' : U\<close> have x'': "x' \<sharp> \<Gamma>"
by (auto dest: valid_typing)
have "(\<lambda>x:T. t) = (\<lambda>x':T'. t')" by fact
also from x x' y have "\<dots> = [(x, y)] \<bullet> [(x', y)] \<bullet> (\<lambda>x':T'. t')"
@@ -592,7 +592,7 @@
then have T: "T = T'" and t: "[(x, y)] \<bullet> [(x', y)] \<bullet> t' = t"
by (simp_all add: trm.inject alpha)
from Abs T have "S = T \<rightarrow> U" by simp
- moreover from `(x', T') # \<Gamma> \<turnstile> t' : U`
+ moreover from \<open>(x', T') # \<Gamma> \<turnstile> t' : U\<close>
have "[(x, y)] \<bullet> [(x', y)] \<bullet> ((x', T') # \<Gamma> \<turnstile> t' : U)"
by (simp add: perm_bool)
with T t y x'' fresh have "(x, T) # \<Gamma> \<turnstile> t : U"
@@ -659,7 +659,7 @@
show ?case
proof (cases q)
case (PVar x' T')
- with `(\<lambda>[PVar x T]. t) = (\<lambda>[q]. u)`
+ with \<open>(\<lambda>[PVar x T]. t) = (\<lambda>[q]. u)\<close>
have "x = x' \<and> t = u \<or> x \<noteq> x' \<and> t = [(x, x')] \<bullet> u \<and> x \<sharp> u"
by (simp add: btrm.inject alpha)
then show ?thesis
@@ -714,7 +714,7 @@
then obtain pi::"name prm" where
"p\<^sub>1 = pi \<bullet> p\<^sub>1'" "(\<lambda>[p\<^sub>2]. t) = pi \<bullet> (\<lambda>[p\<^sub>2']. u)" and
pi: "set pi \<subseteq> (supp p\<^sub>1 \<union> supp p\<^sub>1') \<times> (supp p\<^sub>1 \<union> supp p\<^sub>1')" by auto
- from `(\<lambda>[p\<^sub>2]. t) = pi \<bullet> (\<lambda>[p\<^sub>2']. u)`
+ from \<open>(\<lambda>[p\<^sub>2]. t) = pi \<bullet> (\<lambda>[p\<^sub>2']. u)\<close>
have "(\<lambda>[p\<^sub>2]. t) = (\<lambda>[pi \<bullet> p\<^sub>2']. pi \<bullet> u)"
by (simp add: eqvts)
moreover from PTuple PTuple' have "pat_type p\<^sub>2 = pat_type (pi \<bullet> p\<^sub>2')"
@@ -732,13 +732,13 @@
(supp p\<^sub>2 \<union> supp (pi \<bullet> p\<^sub>2'))" by auto
from PTuple PTuple' have "pi \<bullet> distinct (pat_vars \<langle>\<langle>p\<^sub>1', p\<^sub>2'\<rangle>\<rangle>)" by simp
then have "distinct (pat_vars \<langle>\<langle>pi \<bullet> p\<^sub>1', pi \<bullet> p\<^sub>2'\<rangle>\<rangle>)" by (simp only: eqvts)
- with `p\<^sub>1 = pi \<bullet> p\<^sub>1'` PTuple'
+ with \<open>p\<^sub>1 = pi \<bullet> p\<^sub>1'\<close> PTuple'
have fresh: "(supp p\<^sub>2 \<union> supp (pi \<bullet> p\<^sub>2') :: name set) \<sharp>* p\<^sub>1"
by (auto simp add: set_pat_vars_supp fresh_star_def fresh_def eqvts)
- from `p\<^sub>1 = pi \<bullet> p\<^sub>1'` have "pi' \<bullet> (p\<^sub>1 = pi \<bullet> p\<^sub>1')" by (rule perm_boolI)
+ from \<open>p\<^sub>1 = pi \<bullet> p\<^sub>1'\<close> have "pi' \<bullet> (p\<^sub>1 = pi \<bullet> p\<^sub>1')" by (rule perm_boolI)
with pt_freshs_freshs [OF pt_name_inst at_name_inst pi' fresh fresh]
have "p\<^sub>1 = pi' \<bullet> pi \<bullet> p\<^sub>1'" by (simp add: eqvts)
- with `p\<^sub>2 = pi' \<bullet> pi \<bullet> p\<^sub>2'` have "\<langle>\<langle>p\<^sub>1, p\<^sub>2\<rangle>\<rangle> = (pi' @ pi) \<bullet> \<langle>\<langle>p\<^sub>1', p\<^sub>2'\<rangle>\<rangle>"
+ with \<open>p\<^sub>2 = pi' \<bullet> pi \<bullet> p\<^sub>2'\<close> have "\<langle>\<langle>p\<^sub>1, p\<^sub>2\<rangle>\<rangle> = (pi' @ pi) \<bullet> \<langle>\<langle>p\<^sub>1', p\<^sub>2'\<rangle>\<rangle>"
by (simp add: pt_name2)
moreover
have "((supp p\<^sub>2 \<union> (pi \<bullet> supp p\<^sub>2')) \<times> (supp p\<^sub>2 \<union> (pi \<bullet> supp p\<^sub>2'))::(name \<times> name) set) \<subseteq>
@@ -748,7 +748,7 @@
with pi have "set (pi' @ pi) \<subseteq> (supp \<langle>\<langle>p\<^sub>1, p\<^sub>2\<rangle>\<rangle> \<union> supp \<langle>\<langle>p\<^sub>1', p\<^sub>2'\<rangle>\<rangle>) \<times>
(supp \<langle>\<langle>p\<^sub>1, p\<^sub>2\<rangle>\<rangle> \<union> supp \<langle>\<langle>p\<^sub>1', p\<^sub>2'\<rangle>\<rangle>)"
by (simp add: Sigma_Un_distrib1 Sigma_Un_distrib2 Un_ac) blast
- moreover note `t = pi' \<bullet> pi \<bullet> u`
+ moreover note \<open>t = pi' \<bullet> pi \<bullet> u\<close>
ultimately have "\<langle>\<langle>p\<^sub>1, p\<^sub>2\<rangle>\<rangle> = (pi' @ pi) \<bullet> q \<and> t = (pi' @ pi) \<bullet> u \<and>
set (pi' @ pi) \<subseteq> (supp \<langle>\<langle>p\<^sub>1, p\<^sub>2\<rangle>\<rangle> \<union> supp q) \<times>
(supp \<langle>\<langle>p\<^sub>1, p\<^sub>2\<rangle>\<rangle> \<union> supp q)" using PTuple
@@ -776,10 +776,10 @@
moreover from Let have "pat_type p = pat_type p'"
by (simp add: trm.inject)
moreover note distinct
- moreover from `\<Delta> @ \<Gamma> \<turnstile> u' : U` have "valid (\<Delta> @ \<Gamma>)"
+ moreover from \<open>\<Delta> @ \<Gamma> \<turnstile> u' : U\<close> have "valid (\<Delta> @ \<Gamma>)"
by (rule valid_typing)
then have "valid \<Delta>" by (rule valid_appD)
- with `\<turnstile> p' : T \<Rightarrow> \<Delta>` have "distinct (pat_vars p')"
+ with \<open>\<turnstile> p' : T \<Rightarrow> \<Delta>\<close> have "distinct (pat_vars p')"
by (simp add: valid_distinct pat_vars_ptyping)
ultimately have "\<exists>pi::name prm. p = pi \<bullet> p' \<and> Base u = pi \<bullet> Base u' \<and>
set pi \<subseteq> (supp p \<union> supp p') \<times> (supp p \<union> supp p')"
@@ -788,7 +788,7 @@
and pi': "set pi \<subseteq> (supp p \<union> supp p') \<times> (supp p \<union> supp p')"
by (auto simp add: btrm.inject)
from Let have "\<Gamma> \<turnstile> t : T" by (simp add: trm.inject)
- moreover from `\<turnstile> p' : T \<Rightarrow> \<Delta>` have "\<turnstile> (pi \<bullet> p') : (pi \<bullet> T) \<Rightarrow> (pi \<bullet> \<Delta>)"
+ moreover from \<open>\<turnstile> p' : T \<Rightarrow> \<Delta>\<close> have "\<turnstile> (pi \<bullet> p') : (pi \<bullet> T) \<Rightarrow> (pi \<bullet> \<Delta>)"
by (simp add: ptyping.eqvt)
with pi have "\<turnstile> p : T \<Rightarrow> (pi \<bullet> \<Delta>)" by (simp add: perm_type)
moreover from Let
@@ -805,25 +805,25 @@
shows "\<Gamma> \<turnstile> t' : T" using assms
proof (nominal_induct avoiding: \<Gamma> T rule: eval.strong_induct)
case (TupleL t t' u)
- from `\<Gamma> \<turnstile> \<langle>t, u\<rangle> : T` obtain T\<^sub>1 T\<^sub>2
+ from \<open>\<Gamma> \<turnstile> \<langle>t, u\<rangle> : T\<close> obtain T\<^sub>1 T\<^sub>2
where "T = T\<^sub>1 \<otimes> T\<^sub>2" "\<Gamma> \<turnstile> t : T\<^sub>1" "\<Gamma> \<turnstile> u : T\<^sub>2"
by cases (simp_all add: trm.inject)
- from `\<Gamma> \<turnstile> t : T\<^sub>1` have "\<Gamma> \<turnstile> t' : T\<^sub>1" by (rule TupleL)
- then have "\<Gamma> \<turnstile> \<langle>t', u\<rangle> : T\<^sub>1 \<otimes> T\<^sub>2" using `\<Gamma> \<turnstile> u : T\<^sub>2`
+ from \<open>\<Gamma> \<turnstile> t : T\<^sub>1\<close> have "\<Gamma> \<turnstile> t' : T\<^sub>1" by (rule TupleL)
+ then have "\<Gamma> \<turnstile> \<langle>t', u\<rangle> : T\<^sub>1 \<otimes> T\<^sub>2" using \<open>\<Gamma> \<turnstile> u : T\<^sub>2\<close>
by (rule Tuple)
- with `T = T\<^sub>1 \<otimes> T\<^sub>2` show ?case by simp
+ with \<open>T = T\<^sub>1 \<otimes> T\<^sub>2\<close> show ?case by simp
next
case (TupleR u u' t)
- from `\<Gamma> \<turnstile> \<langle>t, u\<rangle> : T` obtain T\<^sub>1 T\<^sub>2
+ from \<open>\<Gamma> \<turnstile> \<langle>t, u\<rangle> : T\<close> obtain T\<^sub>1 T\<^sub>2
where "T = T\<^sub>1 \<otimes> T\<^sub>2" "\<Gamma> \<turnstile> t : T\<^sub>1" "\<Gamma> \<turnstile> u : T\<^sub>2"
by cases (simp_all add: trm.inject)
- from `\<Gamma> \<turnstile> u : T\<^sub>2` have "\<Gamma> \<turnstile> u' : T\<^sub>2" by (rule TupleR)
- with `\<Gamma> \<turnstile> t : T\<^sub>1` have "\<Gamma> \<turnstile> \<langle>t, u'\<rangle> : T\<^sub>1 \<otimes> T\<^sub>2"
+ from \<open>\<Gamma> \<turnstile> u : T\<^sub>2\<close> have "\<Gamma> \<turnstile> u' : T\<^sub>2" by (rule TupleR)
+ with \<open>\<Gamma> \<turnstile> t : T\<^sub>1\<close> have "\<Gamma> \<turnstile> \<langle>t, u'\<rangle> : T\<^sub>1 \<otimes> T\<^sub>2"
by (rule Tuple)
- with `T = T\<^sub>1 \<otimes> T\<^sub>2` show ?case by simp
+ with \<open>T = T\<^sub>1 \<otimes> T\<^sub>2\<close> show ?case by simp
next
case (Abs t t' x S)
- from `\<Gamma> \<turnstile> (\<lambda>x:S. t) : T` `x \<sharp> \<Gamma>` obtain U where
+ from \<open>\<Gamma> \<turnstile> (\<lambda>x:S. t) : T\<close> \<open>x \<sharp> \<Gamma>\<close> obtain U where
T: "T = S \<rightarrow> U" and U: "(x, S) # \<Gamma> \<turnstile> t : U"
by (rule typing_case_Abs)
from U have "(x, S) # \<Gamma> \<turnstile> t' : U" by (rule Abs)
@@ -832,27 +832,27 @@
with T show ?case by simp
next
case (Beta x u S t)
- from `\<Gamma> \<turnstile> (\<lambda>x:S. t) \<cdot> u : T` `x \<sharp> \<Gamma>`
+ from \<open>\<Gamma> \<turnstile> (\<lambda>x:S. t) \<cdot> u : T\<close> \<open>x \<sharp> \<Gamma>\<close>
obtain "(x, S) # \<Gamma> \<turnstile> t : T" and "\<Gamma> \<turnstile> u : S"
by cases (auto simp add: trm.inject ty.inject elim: typing_case_Abs)
then show ?case by (rule subst_type)
next
case (Let p t \<theta> u)
- from `\<Gamma> \<turnstile> (LET p = t IN u) : T` `supp p \<sharp>* \<Gamma>` `distinct (pat_vars p)`
+ from \<open>\<Gamma> \<turnstile> (LET p = t IN u) : T\<close> \<open>supp p \<sharp>* \<Gamma>\<close> \<open>distinct (pat_vars p)\<close>
obtain U \<Delta> where "\<turnstile> p : U \<Rightarrow> \<Delta>" "\<Gamma> \<turnstile> t : U" "\<Delta> @ \<Gamma> \<turnstile> u : T"
by (rule typing_case_Let)
- then show ?case using `\<turnstile> p \<rhd> t \<Rightarrow> \<theta>` `supp p \<sharp>* t`
+ then show ?case using \<open>\<turnstile> p \<rhd> t \<Rightarrow> \<theta>\<close> \<open>supp p \<sharp>* t\<close>
by (rule match_type)
next
case (AppL t t' u)
- from `\<Gamma> \<turnstile> t \<cdot> u : T` obtain U where
+ from \<open>\<Gamma> \<turnstile> t \<cdot> u : T\<close> obtain U where
t: "\<Gamma> \<turnstile> t : U \<rightarrow> T" and u: "\<Gamma> \<turnstile> u : U"
by cases (auto simp add: trm.inject)
from t have "\<Gamma> \<turnstile> t' : U \<rightarrow> T" by (rule AppL)
then show ?case using u by (rule typing.App)
next
case (AppR u u' t)
- from `\<Gamma> \<turnstile> t \<cdot> u : T` obtain U where
+ from \<open>\<Gamma> \<turnstile> t \<cdot> u : T\<close> obtain U where
t: "\<Gamma> \<turnstile> t : U \<rightarrow> T" and u: "\<Gamma> \<turnstile> u : U"
by cases (auto simp add: trm.inject)
from u have "\<Gamma> \<turnstile> u' : U" by (rule AppR)
--- a/src/HOL/Nominal/Examples/SN.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Nominal/Examples/SN.thy Thu May 26 17:51:22 2016 +0200
@@ -2,9 +2,9 @@
imports Lam_Funs
begin
-text {* Strong Normalisation proof from the Proofs and Types book *}
+text \<open>Strong Normalisation proof from the Proofs and Types book\<close>
-section {* Beta Reduction *}
+section \<open>Beta Reduction\<close>
lemma subst_rename:
assumes a: "c\<sharp>t1"
@@ -110,7 +110,7 @@
by (nominal_induct M M' avoiding: x N rule: Beta.strong_induct)
(auto simp add: fresh_atm subst_lemma fresh_fact)
-section {* types *}
+section \<open>types\<close>
nominal_datatype ty =
TVar "nat"
@@ -156,7 +156,7 @@
nominal_inductive typing
by (simp_all add: abs_fresh fresh_ty)
-subsection {* a fact about beta *}
+subsection \<open>a fact about beta\<close>
definition "NORMAL" :: "lam \<Rightarrow> bool" where
"NORMAL t \<equiv> \<not>(\<exists>t'. t\<longrightarrow>\<^sub>\<beta> t')"
@@ -171,7 +171,7 @@
thus "NORMAL (Var a)" by (auto simp add: NORMAL_def)
qed
-text {* Inductive version of Strong Normalisation *}
+text \<open>Inductive version of Strong Normalisation\<close>
inductive
SN :: "lam \<Rightarrow> bool"
where
@@ -223,7 +223,7 @@
apply(blast)
done
-section {* Candidates *}
+section \<open>Candidates\<close>
nominal_primrec
RED :: "ty \<Rightarrow> lam set"
@@ -232,7 +232,7 @@
| "RED (\<tau>\<rightarrow>\<sigma>) = {t. \<forall>u. (u\<in>RED \<tau> \<longrightarrow> (App t u)\<in>RED \<sigma>)}"
by (rule TrueI)+
-text {* neutral terms *}
+text \<open>neutral terms\<close>
definition NEUT :: "lam \<Rightarrow> bool" where
"NEUT t \<equiv> (\<exists>a. t = Var a) \<or> (\<exists>t1 t2. t = App t1 t2)"
@@ -270,7 +270,7 @@
then show "SN (fst_app (App t s))" by simp
qed
-section {* Candidates *}
+section \<open>Candidates\<close>
definition "CR1" :: "ty \<Rightarrow> bool" where
"CR1 \<tau> \<equiv> \<forall>t. (t\<in>RED \<tau> \<longrightarrow> SN t)"
@@ -338,7 +338,7 @@
ultimately show "App t u \<in> RED \<sigma>" using c3 by (simp add: CR3_def)
qed
-text {* properties of the candiadates *}
+text \<open>properties of the candiadates\<close>
lemma RED_props:
shows "CR1 \<tau>" and "CR2 \<tau>" and "CR3 \<tau>"
proof (nominal_induct \<tau> rule: ty.strong_induct)
@@ -390,10 +390,10 @@
}
qed
-text {*
+text \<open>
the next lemma not as simple as on paper, probably because of
the stronger double_SN induction
-*}
+\<close>
lemma abs_RED:
assumes asm: "\<forall>s\<in>RED \<tau>. t[x::=s]\<in>RED \<sigma>"
shows "Lam [x].t\<in>RED (\<tau>\<rightarrow>\<sigma>)"
@@ -509,7 +509,7 @@
shows "\<theta><t> \<in> RED \<tau>"
using a b
proof(nominal_induct avoiding: \<theta> rule: typing.strong_induct)
- case (t3 a \<Gamma> \<sigma> t \<tau> \<theta>) --"lambda case"
+ case (t3 a \<Gamma> \<sigma> t \<tau> \<theta>) \<comment>"lambda case"
have ih: "\<And>\<theta>. \<theta> closes ((a,\<sigma>)#\<Gamma>) \<Longrightarrow> \<theta><t> \<in> RED \<tau>" by fact
have \<theta>_cond: "\<theta> closes \<Gamma>" by fact
have fresh: "a\<sharp>\<Gamma>" "a\<sharp>\<theta>" by fact+
@@ -519,7 +519,7 @@
then show "\<theta><(Lam [a].t)> \<in> RED (\<sigma> \<rightarrow> \<tau>)" using fresh by simp
qed auto
-section {* identity substitution generated from a context \<Gamma> *}
+section \<open>identity substitution generated from a context \<Gamma>\<close>
fun
"id" :: "(name\<times>ty) list \<Rightarrow> (name\<times>lam) list"
where
@@ -547,12 +547,12 @@
shows "(id \<Gamma>) closes \<Gamma>"
apply(auto)
apply(simp add: id_maps)
-apply(subgoal_tac "CR3 T") --"A"
+apply(subgoal_tac "CR3 T") \<comment>"A"
apply(drule CR3_implies_CR4)
apply(simp add: CR4_def)
apply(drule_tac x="Var x" in spec)
apply(force simp add: NEUT_def NORMAL_Var)
---"A"
+\<comment>"A"
apply(rule RED_props)
done
--- a/src/HOL/Nominal/Examples/SOS.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Nominal/Examples/SOS.thy Thu May 26 17:51:22 2016 +0200
@@ -16,7 +16,7 @@
atom_decl name
-text {* types and terms *}
+text \<open>types and terms\<close>
nominal_datatype ty =
TVar "nat"
| Arrow "ty" "ty" ("_\<rightarrow>_" [100,100] 100)
@@ -33,7 +33,7 @@
by (induct T rule: ty.induct)
(auto simp add: fresh_nat)
-text {* Parallel and single substitution. *}
+text \<open>Parallel and single substitution.\<close>
fun
lookup :: "(name\<times>trm) list \<Rightarrow> name \<Rightarrow> trm"
where
@@ -94,7 +94,7 @@
by (nominal_induct t rule: trm.strong_induct)
(auto simp add: fresh_list_nil)
-text {* Single substitution *}
+text \<open>Single substitution\<close>
abbreviation
subst :: "trm \<Rightarrow> name \<Rightarrow> trm \<Rightarrow> trm" ("_[_::=_]" [100,100,100] 100)
where
@@ -126,7 +126,7 @@
by (nominal_induct e avoiding: \<theta> x e' rule: trm.strong_induct)
(auto simp add: fresh_list_cons fresh_atm forget lookup_fresh lookup_fresh')
-text {* Typing Judgements *}
+text \<open>Typing Judgements\<close>
inductive
valid :: "(name\<times>ty) list \<Rightarrow> bool"
@@ -158,7 +158,7 @@
using a1 a2 a3
by (induct) (auto simp add: fresh_set fresh_prod fresh_atm)
-text {* Typing Relation *}
+text \<open>Typing Relation\<close>
inductive
typing :: "(name\<times>ty) list\<Rightarrow>trm\<Rightarrow>ty\<Rightarrow>bool" ("_ \<turnstile> _ : _" [60,60,60] 60)
@@ -244,7 +244,7 @@
using a b type_substitutivity_aux[where \<Delta>="[]"]
by (auto)
-text {* Values *}
+text \<open>Values\<close>
inductive
val :: "trm\<Rightarrow>bool"
where
@@ -258,7 +258,7 @@
"\<not> val (Var x)"
by (auto elim: val.cases)
-text {* Big-Step Evaluation *}
+text \<open>Big-Step Evaluation\<close>
inductive
big :: "trm\<Rightarrow>trm\<Rightarrow>bool" ("_ \<Down> _" [80,80] 80)
@@ -428,7 +428,7 @@
shows "v \<Down> v"
using a by (simp add: values_reduce_to_themselves Vs_are_values)
-text {* '\<theta> maps x to e' asserts that \<theta> substitutes x with e *}
+text \<open>'\<theta> maps x to e' asserts that \<theta> substitutes x with e\<close>
abbreviation
mapsto :: "(name\<times>trm) list \<Rightarrow> name \<Rightarrow> trm \<Rightarrow> bool" ("_ maps _ to _" [55,55,55] 55)
where
--- a/src/HOL/Nominal/Examples/Standardization.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Nominal/Examples/Standardization.thy Thu May 26 17:51:22 2016 +0200
@@ -3,18 +3,18 @@
Copyright 2005, 2008 TU Muenchen
*)
-section {* Standardization *}
+section \<open>Standardization\<close>
theory Standardization
imports "../Nominal"
begin
-text {*
+text \<open>
The proof of the standardization theorem, as well as most of the theorems about
-lambda calculus in the following sections, are taken from @{text "HOL/Lambda"}.
-*}
+lambda calculus in the following sections, are taken from \<open>HOL/Lambda\<close>.
+\<close>
-subsection {* Lambda terms *}
+subsection \<open>Lambda terms\<close>
atom_decl name
@@ -118,7 +118,7 @@
"s \<rightarrow>\<^sub>\<beta>\<^sup>* t \<equiv> beta\<^sup>*\<^sup>* s t"
-subsection {* Application of a term to a list of terms *}
+subsection \<open>Application of a term to a list of terms\<close>
abbreviation
list_application :: "lam \<Rightarrow> lam list \<Rightarrow> lam" (infixl "\<degree>\<degree>" 150) where
@@ -196,7 +196,7 @@
by (induct ts rule: rev_induct)
(auto simp add: fresh_list_append fresh_list_nil fresh_list_cons)
-text {* A customized induction schema for @{text "\<degree>\<degree>"}. *}
+text \<open>A customized induction schema for \<open>\<degree>\<degree>\<close>.\<close>
lemma lem:
assumes "\<And>n ts (z::'a::fs_name). (\<And>z. \<forall>t \<in> set ts. P z t) \<Longrightarrow> P z (Var n \<degree>\<degree> ts)"
@@ -250,7 +250,7 @@
done
-subsection {* Congruence rules *}
+subsection \<open>Congruence rules\<close>
lemma apps_preserves_beta [simp]:
"r \<rightarrow>\<^sub>\<beta> s \<Longrightarrow> r \<degree>\<degree> ss \<rightarrow>\<^sub>\<beta> s \<degree>\<degree> ss"
@@ -273,7 +273,7 @@
by (blast intro!: rtrancl_beta_AppL rtrancl_beta_AppR intro: rtranclp_trans)
-subsection {* Lifting an order to lists of elements *}
+subsection \<open>Lifting an order to lists of elements\<close>
definition
step1 :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool" where
@@ -339,7 +339,7 @@
done
-subsection {* Lifting beta-reduction to lists *}
+subsection \<open>Lifting beta-reduction to lists\<close>
abbreviation
list_beta :: "lam list \<Rightarrow> lam list \<Rightarrow> bool" (infixl "[\<rightarrow>\<^sub>\<beta>]\<^sub>1" 50) where
@@ -417,12 +417,12 @@
done
-subsection {* Standard reduction relation *}
+subsection \<open>Standard reduction relation\<close>
-text {*
+text \<open>
Based on lecture notes by Ralph Matthes,
original proof idea due to Ralph Loader.
-*}
+\<close>
declare listrel_mono [mono_set]
@@ -514,7 +514,7 @@
from Abs(4) have "ss [\<rightarrow>\<^sub>s] ss'" by (rule listrelp_conj1)
moreover have "[s] [\<rightarrow>\<^sub>s] [s']" by (iprover intro: s listrelp.intros)
ultimately have "ss @ [s] [\<rightarrow>\<^sub>s] ss' @ [s']" by (rule listrelp_app)
- with `r \<rightarrow>\<^sub>s r'` have "(Lam [x].r) \<degree>\<degree> (ss @ [s]) \<rightarrow>\<^sub>s (Lam [x].r') \<degree>\<degree> (ss' @ [s'])"
+ with \<open>r \<rightarrow>\<^sub>s r'\<close> have "(Lam [x].r) \<degree>\<degree> (ss @ [s]) \<rightarrow>\<^sub>s (Lam [x].r') \<degree>\<degree> (ss' @ [s'])"
by (rule better_sred_Abs)
thus ?case by (simp only: app_last)
next
@@ -553,9 +553,9 @@
next
case (Abs y ss ss' r r')
then have "r[x::=s] \<rightarrow>\<^sub>s r'[x::=s']" by fast
- moreover from Abs(8) `s \<rightarrow>\<^sub>s s'` have "map (\<lambda>t. t[x::=s]) ss [\<rightarrow>\<^sub>s] map (\<lambda>t. t[x::=s']) ss'"
+ moreover from Abs(8) \<open>s \<rightarrow>\<^sub>s s'\<close> have "map (\<lambda>t. t[x::=s]) ss [\<rightarrow>\<^sub>s] map (\<lambda>t. t[x::=s']) ss'"
by induct (auto intro: listrelp.intros Abs)
- ultimately show ?case using Abs(6) `y \<sharp> x` `y \<sharp> s` `y \<sharp> s'`
+ ultimately show ?case using Abs(6) \<open>y \<sharp> x\<close> \<open>y \<sharp> s\<close> \<open>y \<sharp> s'\<close>
by simp (rule better_sred_Abs)
next
case (Beta y u ss t r)
@@ -607,10 +607,10 @@
with r'' show ?case by simp
next
case (Abs x ss ss' r r')
- from `(Lam [x].r') \<degree>\<degree> ss' \<rightarrow>\<^sub>\<beta> r''` show ?case
+ from \<open>(Lam [x].r') \<degree>\<degree> ss' \<rightarrow>\<^sub>\<beta> r''\<close> show ?case
proof (cases rule: apps_betasE [where x=x])
case (appL s)
- then obtain r''' where s: "s = (Lam [x].r''')" and r''': "r' \<rightarrow>\<^sub>\<beta> r'''" using `x \<sharp> r''`
+ then obtain r''' where s: "s = (Lam [x].r''')" and r''': "r' \<rightarrow>\<^sub>\<beta> r'''" using \<open>x \<sharp> r''\<close>
by (cases rule: beta.strong_cases) (auto simp add: abs_fresh lam.inject alpha)
from r''' have "r \<rightarrow>\<^sub>s r'''" by (blast intro: Abs)
moreover from Abs have "ss [\<rightarrow>\<^sub>s] ss'" by (iprover dest: listrelp_conj1)
@@ -618,9 +618,9 @@
with appL s show "(Lam [x].r) \<degree>\<degree> ss \<rightarrow>\<^sub>s r''" by simp
next
case (appR rs')
- from Abs(6) [simplified] `ss' [\<rightarrow>\<^sub>\<beta>]\<^sub>1 rs'`
+ from Abs(6) [simplified] \<open>ss' [\<rightarrow>\<^sub>\<beta>]\<^sub>1 rs'\<close>
have "ss [\<rightarrow>\<^sub>s] rs'" by (rule lemma4_aux)
- with `r \<rightarrow>\<^sub>s r'` have "(Lam [x].r) \<degree>\<degree> ss \<rightarrow>\<^sub>s (Lam [x].r') \<degree>\<degree> rs'" by (rule better_sred_Abs)
+ with \<open>r \<rightarrow>\<^sub>s r'\<close> have "(Lam [x].r) \<degree>\<degree> ss \<rightarrow>\<^sub>s (Lam [x].r') \<degree>\<degree> rs'" by (rule better_sred_Abs)
with appR show "(Lam [x].r) \<degree>\<degree> ss \<rightarrow>\<^sub>s r''" by simp
next
case (beta t u' us')
@@ -630,7 +630,7 @@
from Abs(6) ss' obtain u us where
ss: "ss = u # us" and u: "u \<rightarrow>\<^sub>s u'" and us: "us [\<rightarrow>\<^sub>s] us'"
by cases (auto dest!: listrelp_conj1)
- have "r[x::=u] \<rightarrow>\<^sub>s r'[x::=u']" using `r \<rightarrow>\<^sub>s r'` and u by (rule lemma3)
+ have "r[x::=u] \<rightarrow>\<^sub>s r'[x::=u']" using \<open>r \<rightarrow>\<^sub>s r'\<close> and u by (rule lemma3)
with us have "r[x::=u] \<degree>\<degree> us \<rightarrow>\<^sub>s r'[x::=u'] \<degree>\<degree> us'" by (rule lemma1')
hence "(Lam [x].r) \<degree> u \<degree>\<degree> us \<rightarrow>\<^sub>s r'[x::=u'] \<degree>\<degree> us'" by (rule better_sred_Beta)
with ss r'' Lam_eq show "(Lam [x].r) \<degree>\<degree> ss \<rightarrow>\<^sub>s r''" by (simp add: lam.inject alpha)
@@ -647,7 +647,7 @@
by induct (iprover intro: refl_sred lemma4)+
-subsection {* Terms in normal form *}
+subsection \<open>Terms in normal form\<close>
lemma listsp_eqvt [eqvt]:
assumes xs: "listsp p (xs::'a::pt_name list)"
@@ -681,9 +681,9 @@
thus ?thesis by simp
qed
-text {*
+text \<open>
@{term NF} characterizes exactly the terms that are in normal form.
-*}
+\<close>
lemma NF_eq: "NF t = (\<forall>t'. \<not> t \<rightarrow>\<^sub>\<beta> t')"
proof
@@ -740,7 +740,7 @@
qed
-subsection {* Leftmost reduction and weakly normalizing terms *}
+subsection \<open>Leftmost reduction and weakly normalizing terms\<close>
inductive
lred :: "lam \<Rightarrow> lam \<Rightarrow> bool" (infixl "\<rightarrow>\<^sub>l" 50)
@@ -761,13 +761,13 @@
then show ?case by (rule sred.Var)
next
case (Abs r r' x)
- from `r \<rightarrow>\<^sub>s r'`
+ from \<open>r \<rightarrow>\<^sub>s r'\<close>
have "(Lam [x].r) \<degree>\<degree> [] \<rightarrow>\<^sub>s (Lam [x].r') \<degree>\<degree> []" using listrelp.Nil
by (rule better_sred_Abs)
then show ?case by simp
next
case (Beta r x s ss t)
- from `r[x::=s] \<degree>\<degree> ss \<rightarrow>\<^sub>s t`
+ from \<open>r[x::=s] \<degree>\<degree> ss \<rightarrow>\<^sub>s t\<close>
show ?case by (rule better_sred_Beta)
qed
@@ -810,7 +810,7 @@
shows "NF r' \<Longrightarrow> r \<rightarrow>\<^sub>l r'" using r
proof induct
case (Var rs rs' x)
- from `NF (Var x \<degree>\<degree> rs')` have "listsp NF rs'"
+ from \<open>NF (Var x \<degree>\<degree> rs')\<close> have "listsp NF rs'"
by cases simp_all
with Var(1) have "rs [\<rightarrow>\<^sub>l] rs'"
proof induct
@@ -824,7 +824,7 @@
thus ?case by (rule lred.Var)
next
case (Abs x ss ss' r r')
- from `NF ((Lam [x].r') \<degree>\<degree> ss')`
+ from \<open>NF ((Lam [x].r') \<degree>\<degree> ss')\<close>
have ss': "ss' = []" by (rule Abs_NF)
from Abs(4) have ss: "ss = []" using ss'
by cases simp_all
--- a/src/HOL/Nominal/Examples/Support.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Nominal/Examples/Support.thy Thu May 26 17:51:22 2016 +0200
@@ -2,7 +2,7 @@
imports "../Nominal"
begin
-text {*
+text \<open>
An example showing that in general
x\<sharp>(A \<union> B) does not imply x\<sharp>A and x\<sharp>B
@@ -11,31 +11,31 @@
the set of odd atoms. Then A \<union> B, that is the set of
all atoms, has empty support. The sets A, respectively B,
however have the set of all atoms as their support.
-*}
+\<close>
atom_decl atom
-text {* The set of even atoms. *}
+text \<open>The set of even atoms.\<close>
abbreviation
EVEN :: "atom set"
where
"EVEN \<equiv> {atom n | n. \<exists>i. n=2*i}"
-text {* The set of odd atoms: *}
+text \<open>The set of odd atoms:\<close>
abbreviation
ODD :: "atom set"
where
"ODD \<equiv> {atom n | n. \<exists>i. n=2*i+1}"
-text {* An atom is either even or odd. *}
+text \<open>An atom is either even or odd.\<close>
lemma even_or_odd:
fixes n :: nat
shows "\<exists>i. (n = 2*i) \<or> (n=2*i+1)"
by (induct n) (presburger)+
-text {*
+text \<open>
The union of even and odd atoms is the set of all atoms.
- (Unfortunately I do not know a simpler proof of this fact.) *}
+ (Unfortunately I do not know a simpler proof of this fact.)\<close>
lemma EVEN_union_ODD:
shows "EVEN \<union> ODD = UNIV"
using even_or_odd
@@ -49,15 +49,15 @@
finally show "EVEN \<union> ODD = UNIV" by simp
qed
-text {* The sets of even and odd atoms are disjunct. *}
+text \<open>The sets of even and odd atoms are disjunct.\<close>
lemma EVEN_intersect_ODD:
shows "EVEN \<inter> ODD = {}"
using even_or_odd
by (auto) (presburger)
-text {*
+text \<open>
The preceeding two lemmas help us to prove
- the following two useful equalities: *}
+ the following two useful equalities:\<close>
lemma UNIV_subtract:
shows "UNIV - EVEN = ODD"
@@ -65,7 +65,7 @@
using EVEN_union_ODD EVEN_intersect_ODD
by (blast)+
-text {* The sets EVEN and ODD are infinite. *}
+text \<open>The sets EVEN and ODD are infinite.\<close>
lemma EVEN_ODD_infinite:
shows "infinite EVEN"
and "infinite ODD"
@@ -80,10 +80,10 @@
then show "\<exists>f::nat\<Rightarrow>atom. inj f \<and> range f \<subseteq> ODD" by (rule_tac exI)
qed
-text {*
+text \<open>
A general fact about a set S of atoms that is both infinite and
coinfinite. Then S has all atoms as its support. Steve Zdancewic
- helped with proving this fact. *}
+ helped with proving this fact.\<close>
lemma supp_infinite_coinfinite:
fixes S::"atom set"
@@ -114,16 +114,16 @@
then show "(supp S) = (UNIV::atom set)" by auto
qed
-text {* As a corollary we get that EVEN and ODD have infinite support. *}
+text \<open>As a corollary we get that EVEN and ODD have infinite support.\<close>
lemma EVEN_ODD_supp:
shows "supp EVEN = (UNIV::atom set)"
and "supp ODD = (UNIV::atom set)"
using supp_infinite_coinfinite UNIV_subtract EVEN_ODD_infinite
by simp_all
-text {*
+text \<open>
The set of all atoms has empty support, since any swappings leaves
- this set unchanged. *}
+ this set unchanged.\<close>
lemma UNIV_supp:
shows "supp (UNIV::atom set) = ({}::atom set)"
@@ -133,7 +133,7 @@
then show "supp (UNIV::atom set) = ({}::atom set)" by (simp add: supp_def)
qed
-text {* Putting everything together. *}
+text \<open>Putting everything together.\<close>
theorem EVEN_ODD_freshness:
fixes x::"atom"
shows "x\<sharp>(EVEN \<union> ODD)"
@@ -141,6 +141,6 @@
and "\<not>x\<sharp>ODD"
by (auto simp only: fresh_def EVEN_union_ODD EVEN_ODD_supp UNIV_supp)
-text {* Moral: support is a sublte notion. *}
+text \<open>Moral: support is a sublte notion.\<close>
end
--- a/src/HOL/Nominal/Examples/Type_Preservation.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Nominal/Examples/Type_Preservation.thy Thu May 26 17:51:22 2016 +0200
@@ -2,13 +2,13 @@
imports "../Nominal"
begin
-text {*
+text \<open>
This theory shows how to prove the type preservation
property for the simply-typed lambda-calculus and
beta-reduction.
-*}
+\<close>
atom_decl name
@@ -17,7 +17,7 @@
| App "lam" "lam"
| Lam "\<guillemotleft>name\<guillemotright>lam" ("Lam [_]._")
-text {* Capture-Avoiding Substitution *}
+text \<open>Capture-Avoiding Substitution\<close>
nominal_primrec
subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam" ("_[_::=_]")
@@ -43,7 +43,7 @@
by (nominal_induct t avoiding: z y s rule: lam.strong_induct)
(auto simp add: abs_fresh fresh_prod fresh_atm)
-text {* Types *}
+text \<open>Types\<close>
nominal_datatype ty =
TVar "string"
@@ -56,7 +56,7 @@
by (induct T rule: ty.induct)
(auto simp add: fresh_string)
-text {* Typing Contexts *}
+text \<open>Typing Contexts\<close>
type_synonym ctx = "(name\<times>ty) list"
@@ -65,7 +65,7 @@
where
"\<Gamma>\<^sub>1 \<subseteq> \<Gamma>\<^sub>2 \<equiv> \<forall>x. x \<in> set \<Gamma>\<^sub>1 \<longrightarrow> x \<in> set \<Gamma>\<^sub>2"
-text {* Validity of Typing Contexts *}
+text \<open>Validity of Typing Contexts\<close>
inductive
valid :: "(name\<times>ty) list \<Rightarrow> bool"
@@ -99,7 +99,7 @@
using a1 a2 a3
by (induct) (auto simp add: fresh_set fresh_prod fresh_atm)
-text {* Typing Relation *}
+text \<open>Typing Relation\<close>
inductive
typing :: "ctx \<Rightarrow> lam \<Rightarrow> ty \<Rightarrow> bool" ("_ \<turnstile> _ : _")
@@ -166,7 +166,7 @@
using a b
by (auto intro: type_substitution_aux[where \<Delta>="[]",simplified])
-text {* Beta Reduction *}
+text \<open>Beta Reduction\<close>
inductive
"beta" :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^sub>\<beta> _")
--- a/src/HOL/Nominal/Examples/VC_Condition.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Nominal/Examples/VC_Condition.thy Thu May 26 17:51:22 2016 +0200
@@ -2,15 +2,15 @@
imports "../Nominal"
begin
-text {*
+text \<open>
We give here two examples showing that if we use the variable
convention carelessly in rule inductions, we might end
up with faulty lemmas. The point is that the examples
are not variable-convention compatible and therefore in the
nominal data package one is protected from such bogus reasoning.
-*}
+\<close>
-text {* We define alpha-equated lambda-terms as usual. *}
+text \<open>We define alpha-equated lambda-terms as usual.\<close>
atom_decl name
nominal_datatype lam =
@@ -18,13 +18,13 @@
| App "lam" "lam"
| Lam "\<guillemotleft>name\<guillemotright>lam" ("Lam [_]._" [100,100] 100)
-text {*
+text \<open>
The inductive relation 'unbind' unbinds the top-most
binders of a lambda-term; this relation is obviously
not a function, since it does not respect alpha-
equivalence. However as a relation 'unbind' is ok and
a similar relation has been used in our formalisation
- of the algorithm W. *}
+ of the algorithm W.\<close>
inductive
unbind :: "lam \<Rightarrow> name list \<Rightarrow> lam \<Rightarrow> bool" ("_ \<mapsto> _,_" [60,60,60] 60)
@@ -33,10 +33,10 @@
| u_app: "(App t1 t2) \<mapsto> [],(App t1 t2)"
| u_lam: "t\<mapsto>xs,t' \<Longrightarrow> (Lam [x].t) \<mapsto> (x#xs),t'"
-text {*
+text \<open>
We can show that Lam [x].Lam [x].Var x unbinds to [x,x],Var x
and also to [z,y],Var y (though the proof for the second is a
- bit clumsy). *}
+ bit clumsy).\<close>
lemma unbind_lambda_lambda1:
shows "Lam [x].Lam [x].(Var x)\<mapsto>[x,x],(Var x)"
@@ -54,10 +54,10 @@
show "Lam [x].Lam [x].(Var x)\<mapsto>[y,z],(Var z)" by simp
qed
-text {* Unbind is equivariant ...*}
+text \<open>Unbind is equivariant ...\<close>
equivariance unbind
-text {*
+text \<open>
... but it is not variable-convention compatible (see Urban,
Berghofer, Norrish [2007]). This condition requires for rule u_lam to
have the binder x not being a free variable in this rule's conclusion.
@@ -65,15 +65,15 @@
strong induction principle for 'unbind' - that means Isabelle does not
allow us to use the variable convention in induction proofs over 'unbind'.
We can, however, force Isabelle to derive the strengthening induction
- principle and see what happens. *}
+ principle and see what happens.\<close>
nominal_inductive unbind
sorry
-text {*
+text \<open>
To obtain a faulty lemma, we introduce the function 'bind'
which takes a list of names and abstracts them away in
- a given lambda-term. *}
+ a given lambda-term.\<close>
fun
bind :: "name list \<Rightarrow> lam \<Rightarrow> lam"
@@ -81,20 +81,20 @@
"bind [] t = t"
| "bind (x#xs) t = Lam [x].(bind xs t)"
-text {*
+text \<open>
Although not necessary for our main argument below, we can
- easily prove that bind undoes the unbinding. *}
+ easily prove that bind undoes the unbinding.\<close>
lemma bind_unbind:
assumes a: "t \<mapsto> xs,t'"
shows "t = bind xs t'"
using a by (induct) (auto)
-text {*
+text \<open>
The next lemma shows by induction on xs that if x is a free
variable in t, and x does not occur in xs, then x is a free
variable in bind xs t. In the nominal tradition we formulate
- 'is a free variable in' as 'is not fresh for'. *}
+ 'is a free variable in' as 'is not fresh for'.\<close>
lemma free_variable:
fixes x::"name"
@@ -104,13 +104,13 @@
by (induct xs)
(auto simp add: fresh_list_cons abs_fresh fresh_atm)
-text {*
+text \<open>
Now comes the first faulty lemma. It is derived using the
variable convention (i.e. our strong induction principle).
This faulty lemma states that if t unbinds to x#xs and t',
and x is a free variable in t', then it is also a free
variable in bind xs t'. We show this lemma by assuming that
- the binder x is fresh w.r.t. to the xs unbound previously. *}
+ the binder x is fresh w.r.t. to the xs unbound previously.\<close>
lemma faulty1:
assumes a: "t\<mapsto>(x#xs),t'"
@@ -119,10 +119,10 @@
by (nominal_induct t xs'\<equiv>"x#xs" t' avoiding: xs rule: unbind.strong_induct)
(simp_all add: free_variable)
-text {*
+text \<open>
Obviously this lemma is bogus. For example, in
case Lam [x].Lam [x].(Var x) \<mapsto> [x,x],(Var x).
- As a result, we can prove False. *}
+ As a result, we can prove False.\<close>
lemma false1:
shows "False"
@@ -137,10 +137,10 @@
show "False" by simp
qed
-text {*
+text \<open>
The next example is slightly simpler, but looks more
contrived than 'unbind'. This example, called 'strip' just
- strips off the top-most binders from lambdas. *}
+ strips off the top-most binders from lambdas.\<close>
inductive
strip :: "lam \<Rightarrow> lam \<Rightarrow> bool" ("_ \<rightarrow> _" [60,60] 60)
@@ -149,18 +149,18 @@
| s_app: "(App t1 t2) \<rightarrow> (App t1 t2)"
| s_lam: "t \<rightarrow> t' \<Longrightarrow> (Lam [x].t) \<rightarrow> t'"
-text {*
+text \<open>
The relation is equivariant but we have to use again
- sorry to derive a strong induction principle. *}
+ sorry to derive a strong induction principle.\<close>
equivariance strip
nominal_inductive strip
sorry
-text {*
+text \<open>
The second faulty lemma shows that a variable being fresh
- for a term is also fresh for this term after striping. *}
+ for a term is also fresh for this term after striping.\<close>
lemma faulty2:
fixes x::"name"
@@ -170,7 +170,7 @@
by (nominal_induct t t'\<equiv>t' avoiding: t' rule: strip.strong_induct)
(auto simp add: abs_fresh)
-text {* Obviously Lam [x].Var x is a counter example to this lemma. *}
+text \<open>Obviously Lam [x].Var x is a counter example to this lemma.\<close>
lemma false2:
shows "False"
@@ -183,8 +183,8 @@
then show "False" by (simp add: fresh_atm)
qed
-text {* A similar effect can be achieved by using naive inversion
- on rules. *}
+text \<open>A similar effect can be achieved by using naive inversion
+ on rules.\<close>
lemma false3:
shows "False"
@@ -202,9 +202,9 @@
by (cases) (simp_all add: lam.inject fresh_atm)
qed
-text {*
+text \<open>
Moral: Who would have thought that the variable convention
is in general an unsound reasoning principle.
- *}
+\<close>
end
--- a/src/HOL/Nominal/Examples/W.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Nominal/Examples/W.thy Thu May 26 17:51:22 2016 +0200
@@ -2,7 +2,7 @@
imports "../Nominal"
begin
-text {* Example for strong induction rules avoiding sets of atoms. *}
+text \<open>Example for strong induction rules avoiding sets of atoms.\<close>
atom_decl tvar var
@@ -56,7 +56,7 @@
type_synonym
Ctxt = "(var\<times>tyS) list"
-text {* free type variables *}
+text \<open>free type variables\<close>
consts ftv :: "'a \<Rightarrow> tvar list"
@@ -140,7 +140,7 @@
shows "pi\<bullet>(ftv \<Gamma>) = ftv (pi\<bullet>\<Gamma>)"
by (induct \<Gamma>) (auto simp add: eqvts)
-text {* Valid *}
+text \<open>Valid\<close>
inductive
valid :: "Ctxt \<Rightarrow> bool"
where
@@ -149,7 +149,7 @@
equivariance valid
-text {* General *}
+text \<open>General\<close>
primrec gen :: "ty \<Rightarrow> tvar list \<Rightarrow> tyS" where
"gen T [] = Ty T"
| "gen T (X#Xs) = \<forall>[X].(gen T Xs)"
@@ -171,7 +171,7 @@
shows "pi\<bullet>(close \<Gamma> T) = close (pi\<bullet>\<Gamma>) (pi\<bullet>T)"
by (simp_all only: eqvts)
-text {* Substitution *}
+text \<open>Substitution\<close>
type_synonym Subst = "(tvar\<times>ty) list"
@@ -301,7 +301,7 @@
by (nominal_induct T rule: ty.strong_induct)
(auto simp add: fresh_atm)
-text {* instance of a type scheme *}
+text \<open>instance of a type scheme\<close>
inductive
inst :: "ty \<Rightarrow> tyS \<Rightarrow> bool"("_ \<prec> _" [50,51] 50)
where
@@ -352,7 +352,7 @@
done
-text{* typing judgements *}
+text\<open>typing judgements\<close>
inductive
typing :: "Ctxt \<Rightarrow> trm \<Rightarrow> ty \<Rightarrow> bool" (" _ \<turnstile> _ : _ " [60,60,60] 60)
where
--- a/src/HOL/Nominal/Examples/Weakening.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Nominal/Examples/Weakening.thy Thu May 26 17:51:22 2016 +0200
@@ -2,14 +2,14 @@
imports "../Nominal"
begin
-text {*
+text \<open>
A simple proof of the Weakening Property in the simply-typed
lambda-calculus. The proof is simple, because we can make use
- of the variable convention. *}
+ of the variable convention.\<close>
atom_decl name
-text {* Terms and Types *}
+text \<open>Terms and Types\<close>
nominal_datatype lam =
Var "name"
@@ -27,9 +27,9 @@
by (nominal_induct T rule: ty.strong_induct)
(auto simp add: fresh_string)
-text {*
+text \<open>
Valid contexts (at the moment we have no type for finite
- sets yet, therefore typing-contexts are lists). *}
+ sets yet, therefore typing-contexts are lists).\<close>
inductive
valid :: "(name\<times>ty) list \<Rightarrow> bool"
@@ -39,7 +39,7 @@
equivariance valid
-text{* Typing judgements *}
+text\<open>Typing judgements\<close>
inductive
typing :: "(name\<times>ty) list\<Rightarrow>lam\<Rightarrow>ty\<Rightarrow>bool" ("_ \<turnstile> _ : _" [60,60,60] 60)
@@ -48,27 +48,27 @@
| t_App[intro]: "\<lbrakk>\<Gamma> \<turnstile> t1 : T1\<rightarrow>T2; \<Gamma> \<turnstile> t2 : T1\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> App t1 t2 : T2"
| t_Lam[intro]: "\<lbrakk>x\<sharp>\<Gamma>;(x,T1)#\<Gamma> \<turnstile> t : T2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam [x].t : T1\<rightarrow>T2"
-text {*
+text \<open>
We derive the strong induction principle for the typing
relation (this induction principle has the variable convention
- already built-in). *}
+ already built-in).\<close>
equivariance typing
nominal_inductive typing
by (simp_all add: abs_fresh ty_fresh)
-text {* Abbreviation for the notion of subcontexts. *}
+text \<open>Abbreviation for the notion of subcontexts.\<close>
abbreviation
"sub_context" :: "(name\<times>ty) list \<Rightarrow> (name\<times>ty) list \<Rightarrow> bool" ("_ \<subseteq> _" [60,60] 60)
where
"\<Gamma>1 \<subseteq> \<Gamma>2 \<equiv> \<forall>x T. (x,T)\<in>set \<Gamma>1 \<longrightarrow> (x,T)\<in>set \<Gamma>2"
-text {* Now it comes: The Weakening Lemma *}
+text \<open>Now it comes: The Weakening Lemma\<close>
-text {*
+text \<open>
The first version is, after setting up the induction,
- completely automatic except for use of atomize. *}
+ completely automatic except for use of atomize.\<close>
lemma weakening_version1:
fixes \<Gamma>1 \<Gamma>2::"(name\<times>ty) list"
@@ -80,9 +80,9 @@
by (nominal_induct \<Gamma>1 t T avoiding: \<Gamma>2 rule: typing.strong_induct)
(auto | atomize)+
-text {*
+text \<open>
The second version gives the details for the variable
- and lambda case. *}
+ and lambda case.\<close>
lemma weakening_version2:
fixes \<Gamma>1 \<Gamma>2::"(name\<times>ty) list"
@@ -114,10 +114,10 @@
with vc show "\<Gamma>2 \<turnstile> Lam [x].t : T1\<rightarrow>T2" by auto
qed (auto) (* app case *)
-text{*
+text\<open>
The original induction principle for the typing relation
is not strong enough - even this simple lemma fails to be
- simple ;o) *}
+ simple ;o)\<close>
lemma weakening_not_straigh_forward:
fixes \<Gamma>1 \<Gamma>2::"(name\<times>ty) list"
@@ -147,13 +147,13 @@
have "valid ((x,T1)#\<Gamma>2)" using v2 (* fails *)
oops
-text{*
+text\<open>
To show that the proof with explicit renaming is not simple,
here is the proof without using the variable convention. It
crucially depends on the equivariance property of the typing
relation.
-*}
+\<close>
lemma weakening_with_explicit_renaming:
fixes \<Gamma>1 \<Gamma>2::"(name\<times>ty) list"
@@ -199,9 +199,9 @@
ultimately show "\<Gamma>2 \<turnstile> Lam [x].t : T1 \<rightarrow> T2" by simp
qed (auto) (* var and app cases *)
-text {*
+text \<open>
Moral: compare the proof with explicit renamings to weakening_version1
and weakening_version2, and imagine you are proving something more substantial
- than the weakening lemma. *}
+ than the weakening lemma.\<close>
end
--- a/src/HOL/Nominal/Nominal.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Nominal/Nominal.thy Thu May 26 17:51:22 2016 +0200
@@ -9,7 +9,7 @@
declare [[typedef_overloaded]]
-section {* Permutations *}
+section \<open>Permutations\<close>
(*======================*)
type_synonym
@@ -184,13 +184,13 @@
by (induct s)(auto simp add: perm_char_def)
-section {* permutation equality *}
+section \<open>permutation equality\<close>
(*==============================*)
definition prm_eq :: "'x prm \<Rightarrow> 'x prm \<Rightarrow> bool" (" _ \<triangleq> _ " [80,80] 80) where
"pi1 \<triangleq> pi2 \<longleftrightarrow> (\<forall>a::'x. pi1\<bullet>a = pi2\<bullet>a)"
-section {* Support, Freshness and Supports*}
+section \<open>Support, Freshness and Supports\<close>
(*========================================*)
definition supp :: "'a \<Rightarrow> ('x set)" where
"supp x = {a . (infinite {b . [(a,b)]\<bullet>x \<noteq> x})}"
@@ -365,7 +365,7 @@
shows "a\<sharp>b"
by (simp add: fresh_def supp_bool)
-text {* Normalization of freshness results; cf.\ @{text nominal_induct} *}
+text \<open>Normalization of freshness results; cf.\ \<open>nominal_induct\<close>\<close>
lemma fresh_unit_elim:
shows "(a\<sharp>() \<Longrightarrow> PROP C) \<equiv> PROP C"
by (simp add: fresh_def supp_unit)
@@ -385,14 +385,14 @@
and "a\<sharp>(x,y) \<Longrightarrow> a\<sharp>y"
by (simp_all add: fresh_prod)
-ML {*
+ML \<open>
val mksimps_pairs = (@{const_name Nominal.fresh}, @{thms fresh_prodD}) :: mksimps_pairs;
-*}
-declaration {* fn _ =>
+\<close>
+declaration \<open>fn _ =>
Simplifier.map_ss (Simplifier.set_mksimps (mksimps mksimps_pairs))
-*}
-
-section {* Abstract Properties for Permutations and Atoms *}
+\<close>
+
+section \<open>Abstract Properties for Permutations and Atoms\<close>
(*=========================================================*)
(* properties for being a permutation type *)
@@ -425,7 +425,7 @@
definition
"fs TYPE('a) TYPE('x) \<equiv> \<forall>(x::'a). finite ((supp x)::'x set)"
-section {* Lemmas about the atom-type properties*}
+section \<open>Lemmas about the atom-type properties\<close>
(*==============================================*)
lemma at1:
@@ -770,7 +770,7 @@
apply(rule at_prm_eq_refl)
done
---"there always exists an atom that is not being in a finite set"
+\<comment>"there always exists an atom that is not being in a finite set"
lemma ex_in_inf:
fixes A::"'x set"
assumes at: "at TYPE('x)"
@@ -785,7 +785,7 @@
then show ?thesis ..
qed
-text {* there always exists a fresh name for an object with finite support *}
+text \<open>there always exists a fresh name for an object with finite support\<close>
lemma at_exists_fresh':
fixes x :: "'a"
assumes at: "at TYPE('x)"
@@ -833,7 +833,7 @@
then show "\<exists>(b::'x). a\<noteq>b" by blast
qed
---"the at-props imply the pt-props"
+\<comment>"the at-props imply the pt-props"
lemma at_pt_inst:
assumes at: "at TYPE('x)"
shows "pt TYPE('x) TYPE('x)"
@@ -843,7 +843,7 @@
apply(simp only: prm_eq_def)
done
-section {* finite support properties *}
+section \<open>finite support properties\<close>
(*===================================*)
lemma fs1:
@@ -907,7 +907,7 @@
apply(rule fs1[OF fs])
done
-section {* Lemmas about the permutation properties *}
+section \<open>Lemmas about the permutation properties\<close>
(*=================================================*)
lemma pt1:
@@ -941,7 +941,7 @@
shows "pi1 \<triangleq> pi2 \<Longrightarrow> (rev pi1)\<bullet>x = (rev pi2)\<bullet>x"
by (rule pt3[OF pt], simp add: at_prm_rev_eq[OF at])
-section {* composition properties *}
+section \<open>composition properties\<close>
(* ============================== *)
lemma cp1:
fixes pi1::"'x prm"
@@ -960,7 +960,7 @@
apply(rule at_ds8[OF at])
done
-section {* disjointness properties *}
+section \<open>disjointness properties\<close>
(*=================================*)
lemma dj_perm_forget:
fixes pi::"'y prm"
@@ -1006,7 +1006,7 @@
shows "a\<sharp>b"
by (simp add: fresh_def dj_supp[OF dj])
-section {* permutation type instances *}
+section \<open>permutation type instances\<close>
(* ===================================*)
lemma pt_fun_inst:
@@ -1128,7 +1128,7 @@
apply(simp add: pt3[OF pta] pt3[OF ptb])
done
-section {* further lemmas for permutation types *}
+section \<open>further lemmas for permutation types\<close>
(*==============================================*)
lemma pt_rev_pi:
@@ -1354,7 +1354,7 @@
apply(simp add: pt_pi_rev[OF pt, OF at])
done
--- "some helper lemmas for the pt_perm_supp_ineq lemma"
+\<comment> "some helper lemmas for the pt_perm_supp_ineq lemma"
lemma Collect_permI:
fixes pi :: "'x prm"
and x :: "'a"
@@ -1672,7 +1672,7 @@
shows "pi\<bullet>(c\<sharp>x) = (pi\<bullet>c)\<sharp>(pi\<bullet>x)"
by (simp add: pt_fresh_left_ineq[OF pta, OF ptb, OF at, OF cp] dj_perm_forget[OF dj] perm_bool)
---"the co-set of a finite set is infinte"
+\<comment>"the co-set of a finite set is infinte"
lemma finite_infinite:
assumes a: "finite {b::'x. P b}"
and b: "infinite (UNIV::'x set)"
@@ -1787,7 +1787,7 @@
thus ?thesis by (simp add: pt2[OF pt])
qed
-section {* equivariance for some connectives *}
+section \<open>equivariance for some connectives\<close>
lemma pt_all_eqvt:
fixes pi :: "'x prm"
and x :: "'a"
@@ -1833,7 +1833,7 @@
apply(rule theI'[OF unique])
done
-section {* facts about supports *}
+section \<open>facts about supports\<close>
(*==============================*)
lemma supports_subset:
@@ -2002,7 +2002,7 @@
by (simp add: at_fin_set_supp fresh_def at fs)
-section {* Permutations acting on Functions *}
+section \<open>Permutations acting on Functions\<close>
(*==========================================*)
lemma pt_fun_app_eq:
@@ -2015,7 +2015,7 @@
by (simp add: perm_fun_def pt_rev_pi[OF pt, OF at])
---"sometimes pt_fun_app_eq does too much; this lemma 'corrects it'"
+\<comment>"sometimes pt_fun_app_eq does too much; this lemma 'corrects it'"
lemma pt_perm:
fixes x :: "'a"
and pi1 :: "'x prm"
@@ -2056,7 +2056,7 @@
qed
qed
--- "two helper lemmas for the equivariance of functions"
+\<comment> "two helper lemmas for the equivariance of functions"
lemma pt_swap_eq_aux:
fixes y :: "'a"
and pi :: "'x prm"
@@ -2188,7 +2188,7 @@
from a1 a2 a3 show False by (force dest: finite_subset)
qed
-section {* Facts about the support of finite sets of finitely supported things *}
+section \<open>Facts about the support of finite sets of finitely supported things\<close>
(*=============================================================================*)
definition X_to_Un_supp :: "('a set) \<Rightarrow> 'x set" where
@@ -2410,7 +2410,7 @@
by (simp add: fresh_def pt_list_set_supp[OF pt, OF at, OF fs])
-section {* generalisation of freshness to lists and sets of atoms *}
+section \<open>generalisation of freshness to lists and sets of atoms\<close>
(*================================================================*)
consts
@@ -2459,7 +2459,7 @@
"({} \<sharp>* c \<Longrightarrow> PROP C) \<equiv> PROP C"
by (simp add: fresh_star_def)
-text {* Normalization of freshness results; see \ @{text nominal_induct} *}
+text \<open>Normalization of freshness results; see \ \<open>nominal_induct\<close>\<close>
lemma fresh_star_unit_elim:
shows "((a::'a set)\<sharp>*() \<Longrightarrow> PROP C) \<equiv> PROP C"
@@ -2591,14 +2591,14 @@
apply(simp add: pt2[OF pt])
done
-section {* Infrastructure lemmas for strong rule inductions *}
+section \<open>Infrastructure lemmas for strong rule inductions\<close>
(*==========================================================*)
-text {*
+text \<open>
For every set of atoms, there is another set of atoms
avoiding a finitely supported c and there is a permutation
which 'translates' between both sets.
-*}
+\<close>
lemma at_set_avoiding_aux:
fixes Xs::"'a set"
@@ -2676,7 +2676,7 @@
using a b at_set_avoiding_aux[OF at, where Xs="Xs" and As="Xs" and c="c"]
by (blast)
-section {* composition instances *}
+section \<open>composition instances\<close>
(* ============================= *)
lemma cp_list_inst:
@@ -2754,7 +2754,7 @@
done
-section {* Andy's freshness lemma *}
+section \<open>Andy's freshness lemma\<close>
(*================================*)
lemma freshness_lemma:
@@ -2819,7 +2819,7 @@
thus "fr1 = fr2" by force
qed
--- "packaging the freshness lemma into a function"
+\<comment> "packaging the freshness lemma into a function"
definition fresh_fun :: "('x\<Rightarrow>'a)\<Rightarrow>'a" where
"fresh_fun (h) \<equiv> THE fr. (\<forall>(a::'x). a\<sharp>h \<longrightarrow> (h a) = fr)"
@@ -2942,7 +2942,7 @@
apply(simp add: pt_fresh_fresh[OF pt_fun_inst[OF at_pt_inst[OF at], OF pt], OF at, OF at])
done
-section {* Abstraction function *}
+section \<open>Abstraction function\<close>
(*==============================*)
lemma pt_abs_fun_inst:
@@ -3068,7 +3068,7 @@
show "?LHS=?RHS"
proof -
have "(c=a) \<or> (c=b) \<or> (c\<noteq>a \<and> c\<noteq>b)" by blast
- moreover --"case c=a"
+ moreover \<comment>"case c=a"
{ have "nSome(x) = nSome([(a,b)]\<bullet>y)" using a2 by simp
also have "\<dots> = nSome([(b,a)]\<bullet>y)" by (simp, rule pt3[OF pt], rule at_ds5[OF at])
finally have "nSome(x) = nSome([(b,a)]\<bullet>y)" by simp
@@ -3076,7 +3076,7 @@
assume "c=a"
ultimately have "?LHS=?RHS" using a1 a3 by simp
}
- moreover -- "case c=b"
+ moreover \<comment> "case c=b"
{ have a4: "y=[(a,b)]\<bullet>x" using a2 by (simp only: pt_swap_bij[OF pt, OF at])
hence "a\<sharp>([(a,b)]\<bullet>x)" using a3 by simp
hence "b\<sharp>x" by (simp add: at_calc[OF at] pt_fresh_left[OF pt, OF at])
@@ -3084,7 +3084,7 @@
assume "c=b"
ultimately have "?LHS=?RHS" using a1 a4 by simp
}
- moreover -- "case c\<noteq>a \<and> c\<noteq>b"
+ moreover \<comment> "case c\<noteq>a \<and> c\<noteq>b"
{ assume a5: "c\<noteq>a \<and> c\<noteq>b"
moreover
have "c\<sharp>x = c\<sharp>y" using a2 a5 by (force simp add: at_calc[OF at] pt_fresh_left[OF pt, OF at])
@@ -3376,7 +3376,7 @@
shows "b\<sharp>([a].x) = b\<sharp>x"
by (simp add: fresh_def abs_fun_supp_ineq[OF pta, OF ptb, OF at, OF cp, OF dj])
-section {* abstraction type for the parsing in nominal datatype *}
+section \<open>abstraction type for the parsing in nominal datatype\<close>
(*==============================================================*)
inductive_set ABS_set :: "('x\<Rightarrow>('a noption)) set"
@@ -3395,7 +3395,7 @@
qed
-section {* lemmas for deciding permutation equations *}
+section \<open>lemmas for deciding permutation equations\<close>
(*===================================================*)
lemma perm_aux_fold:
@@ -3438,7 +3438,7 @@
shows "((pi\<bullet>(\<lambda>x. f x))=y) = ((\<lambda>x. (pi\<bullet>(f ((rev pi)\<bullet>x))))=y)"
by (simp add: perm_fun_def)
-section {* test *}
+section \<open>test\<close>
lemma at_prm_eq_compose:
fixes pi1 :: "'x prm"
and pi2 :: "'x prm"
@@ -3604,58 +3604,58 @@
ML_file "nominal_permeq.ML"
method_setup perm_simp =
- {* NominalPermeq.perm_simp_meth *}
- {* simp rules and simprocs for analysing permutations *}
+ \<open>NominalPermeq.perm_simp_meth\<close>
+ \<open>simp rules and simprocs for analysing permutations\<close>
method_setup perm_simp_debug =
- {* NominalPermeq.perm_simp_meth_debug *}
- {* simp rules and simprocs for analysing permutations including debugging facilities *}
+ \<open>NominalPermeq.perm_simp_meth_debug\<close>
+ \<open>simp rules and simprocs for analysing permutations including debugging facilities\<close>
method_setup perm_extend_simp =
- {* NominalPermeq.perm_extend_simp_meth *}
- {* tactic for deciding equalities involving permutations *}
+ \<open>NominalPermeq.perm_extend_simp_meth\<close>
+ \<open>tactic for deciding equalities involving permutations\<close>
method_setup perm_extend_simp_debug =
- {* NominalPermeq.perm_extend_simp_meth_debug *}
- {* tactic for deciding equalities involving permutations including debugging facilities *}
+ \<open>NominalPermeq.perm_extend_simp_meth_debug\<close>
+ \<open>tactic for deciding equalities involving permutations including debugging facilities\<close>
method_setup supports_simp =
- {* NominalPermeq.supports_meth *}
- {* tactic for deciding whether something supports something else *}
+ \<open>NominalPermeq.supports_meth\<close>
+ \<open>tactic for deciding whether something supports something else\<close>
method_setup supports_simp_debug =
- {* NominalPermeq.supports_meth_debug *}
- {* tactic for deciding whether something supports something else including debugging facilities *}
+ \<open>NominalPermeq.supports_meth_debug\<close>
+ \<open>tactic for deciding whether something supports something else including debugging facilities\<close>
method_setup finite_guess =
- {* NominalPermeq.finite_guess_meth *}
- {* tactic for deciding whether something has finite support *}
+ \<open>NominalPermeq.finite_guess_meth\<close>
+ \<open>tactic for deciding whether something has finite support\<close>
method_setup finite_guess_debug =
- {* NominalPermeq.finite_guess_meth_debug *}
- {* tactic for deciding whether something has finite support including debugging facilities *}
+ \<open>NominalPermeq.finite_guess_meth_debug\<close>
+ \<open>tactic for deciding whether something has finite support including debugging facilities\<close>
method_setup fresh_guess =
- {* NominalPermeq.fresh_guess_meth *}
- {* tactic for deciding whether an atom is fresh for something*}
+ \<open>NominalPermeq.fresh_guess_meth\<close>
+ \<open>tactic for deciding whether an atom is fresh for something\<close>
method_setup fresh_guess_debug =
- {* NominalPermeq.fresh_guess_meth_debug *}
- {* tactic for deciding whether an atom is fresh for something including debugging facilities *}
+ \<open>NominalPermeq.fresh_guess_meth_debug\<close>
+ \<open>tactic for deciding whether an atom is fresh for something including debugging facilities\<close>
(*****************************************************************)
(* tactics for generating fresh names and simplifying fresh_funs *)
ML_file "nominal_fresh_fun.ML"
-method_setup generate_fresh = {*
+method_setup generate_fresh = \<open>
Args.type_name {proper = true, strict = true} >>
(fn s => fn ctxt => SIMPLE_METHOD (generate_fresh_tac ctxt s))
-*} "generate a name fresh for all the variables in the goal"
-
-method_setup fresh_fun_simp = {*
+\<close> "generate a name fresh for all the variables in the goal"
+
+method_setup fresh_fun_simp = \<open>
Scan.lift (Args.parens (Args.$$$ "no_asm") >> K true || Scan.succeed false) >>
(fn b => fn ctxt => SIMPLE_METHOD' (fresh_fun_tac ctxt b))
-*} "delete one inner occurrence of fresh_fun"
+\<close> "delete one inner occurrence of fresh_fun"
(************************************************)
@@ -3678,7 +3678,7 @@
(* setup for induction principles method *)
ML_file "nominal_induct.ML"
method_setup nominal_induct =
- {* NominalInduct.nominal_induct_method *}
- {* nominal induction *}
+ \<open>NominalInduct.nominal_induct_method\<close>
+ \<open>nominal induction\<close>
end
--- a/src/HOL/Number_Theory/Cong.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Number_Theory/Cong.thy Thu May 26 17:51:22 2016 +0200
@@ -29,7 +29,7 @@
imports Primes
begin
-subsection \<open>Turn off @{text One_nat_def}\<close>
+subsection \<open>Turn off \<open>One_nat_def\<close>\<close>
lemma power_eq_one_eq_nat [simp]: "((x::nat)^m = 1) = (m = 0 | x = 1)"
by (induct m) auto
--- a/src/HOL/Number_Theory/Euclidean_Algorithm.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Number_Theory/Euclidean_Algorithm.thy Thu May 26 17:51:22 2016 +0200
@@ -85,7 +85,7 @@
where
"lcm_eucl a b = normalize (a * b) div gcd_eucl a b"
-definition Lcm_eucl :: "'a set \<Rightarrow> 'a" -- \<open>
+definition Lcm_eucl :: "'a set \<Rightarrow> 'a" \<comment> \<open>
Somewhat complicated definition of Lcm that has the advantage of working
for infinite sets as well\<close>
where
--- a/src/HOL/Number_Theory/Fib.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Number_Theory/Fib.thy Thu May 26 17:51:22 2016 +0200
@@ -97,7 +97,7 @@
qed
lemma fib_gcd: "fib (gcd m n) = gcd (fib m) (fib n)"
- -- \<open>Law 6.111\<close>
+ \<comment> \<open>Law 6.111\<close>
by (induct m n rule: gcd_nat_induct) (simp_all add: gcd_non_0_nat gcd.commute gcd_fib_mod)
theorem fib_mult_eq_setsum_nat: "fib (Suc n) * fib n = (\<Sum>k \<in> {..n}. fib k * fib k)"
--- a/src/HOL/Number_Theory/Residues.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Number_Theory/Residues.thy Thu May 26 17:51:22 2016 +0200
@@ -124,7 +124,7 @@
by (subst res_units_eq) auto
text \<open>
- The function @{text "a \<mapsto> a mod m"} maps the integers to the
+ The function \<open>a \<mapsto> a mod m\<close> maps the integers to the
residue classes. The following lemmas show that this mapping
respects addition and multiplication on the integers.
\<close>
@@ -327,7 +327,7 @@
*)
-text \<open>Outside the locale, we can relax the restriction @{text "m > 1"}.\<close>
+text \<open>Outside the locale, we can relax the restriction \<open>m > 1\<close>.\<close>
lemma euler_theorem:
assumes "m \<ge> 0"
and "gcd a m = 1"
--- a/src/HOL/Old_Number_Theory/Chinese.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Old_Number_Theory/Chinese.thy Thu May 26 17:51:22 2016 +0200
@@ -11,8 +11,7 @@
text \<open>
The Chinese Remainder Theorem for an arbitrary finite number of
- equations. (The one-equation case is included in theory @{text
- IntPrimes}. Uses functions for indexing.\footnote{Maybe @{term
+ equations. (The one-equation case is included in theory \<open>IntPrimes\<close>. Uses functions for indexing.\footnote{Maybe @{term
funprod} and @{term funsum} should be based on general @{term fold}
on indices?}
\<close>
--- a/src/HOL/Old_Number_Theory/EulerFermat.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Old_Number_Theory/EulerFermat.thy Thu May 26 17:51:22 2016 +0200
@@ -54,11 +54,11 @@
where "zcongm m = (\<lambda>a b. zcong a b m)"
lemma abs_eq_1_iff [iff]: "(\<bar>z\<bar> = (1::int)) = (z = 1 \<or> z = -1)"
- -- \<open>LCP: not sure why this lemma is needed now\<close>
+ \<comment> \<open>LCP: not sure why this lemma is needed now\<close>
by (auto simp add: abs_if)
-text \<open>\medskip @{text norRRset}\<close>
+text \<open>\medskip \<open>norRRset\<close>\<close>
declare BnorRset.simps [simp del]
@@ -249,7 +249,7 @@
\<Prod>(BnorRset a m) * x^card (BnorRset a m)"
apply (induct a m rule: BnorRset_induct)
prefer 2
- apply (simplesubst BnorRset.simps) --\<open>multiple redexes\<close>
+ apply (simplesubst BnorRset.simps) \<comment>\<open>multiple redexes\<close>
apply (unfold Let_def, auto)
apply (simp add: Bnor_fin Bnor_mem_zle_swap)
apply (subst setprod.insert)
--- a/src/HOL/Old_Number_Theory/Fib.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Old_Number_Theory/Fib.thy Thu May 26 17:51:22 2016 +0200
@@ -30,7 +30,7 @@
to the Simplifier and are applied very selectively at first.
\<close>
-text\<open>We disable @{text fib.fib_2fib_2} for simplification ...\<close>
+text\<open>We disable \<open>fib.fib_2fib_2\<close> for simplification ...\<close>
declare fib_2 [simp del]
text\<open>...then prove a version that has a more restrictive pattern.\<close>
@@ -133,7 +133,7 @@
qed
qed
-lemma fib_gcd: "fib (gcd m n) = gcd (fib m) (fib n)" -- \<open>Law 6.111\<close>
+lemma fib_gcd: "fib (gcd m n) = gcd (fib m) (fib n)" \<comment> \<open>Law 6.111\<close>
apply (induct m n rule: gcd_induct)
apply (simp_all add: gcd_non_0 gcd_commute gcd_fib_mod)
done
--- a/src/HOL/Old_Number_Theory/IntFact.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Old_Number_Theory/IntFact.thy Thu May 26 17:51:22 2016 +0200
@@ -11,7 +11,7 @@
text \<open>
Factorial on integers and recursively defined set including all
- Integers from @{text 2} up to @{text a}. Plus definition of product
+ Integers from \<open>2\<close> up to \<open>a\<close>. Plus definition of product
of finite set.
\bigskip
@@ -26,7 +26,7 @@
text \<open>
\medskip @{term d22set} --- recursively defined set including all
- integers from @{text 2} up to @{text a}
+ integers from \<open>2\<close> up to \<open>a\<close>
\<close>
declare d22set.simps [simp del]
--- a/src/HOL/Old_Number_Theory/IntPrimes.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Old_Number_Theory/IntPrimes.thy Thu May 26 17:51:22 2016 +0200
@@ -10,11 +10,10 @@
begin
text \<open>
- The @{text dvd} relation, GCD, Euclid's extended algorithm, primes,
- congruences (all on the Integers). Comparable to theory @{text
- Primes}, but @{text dvd} is included here as it is not present in
+ The \<open>dvd\<close> relation, GCD, Euclid's extended algorithm, primes,
+ congruences (all on the Integers). Comparable to theory \<open>Primes\<close>, but \<open>dvd\<close> is included here as it is not present in
main HOL. Also includes extended GCD and congruences not present in
- @{text Primes}.
+ \<open>Primes\<close>.
\<close>
--- a/src/HOL/Old_Number_Theory/Legacy_GCD.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Old_Number_Theory/Legacy_GCD.thy Thu May 26 17:51:22 2016 +0200
@@ -16,7 +16,7 @@
subsection \<open>Specification of GCD on nats\<close>
definition
- is_gcd :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" where -- \<open>@{term gcd} as a relation\<close>
+ is_gcd :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" where \<comment> \<open>@{term gcd} as a relation\<close>
"is_gcd m n p \<longleftrightarrow> p dvd m \<and> p dvd n \<and>
(\<forall>d. d dvd m \<longrightarrow> d dvd n \<longrightarrow> d dvd p)"
@@ -69,7 +69,7 @@
declare gcd.simps [simp del]
text \<open>
- \medskip @{term "gcd m n"} divides @{text m} and @{text n}. The
+ \medskip @{term "gcd m n"} divides \<open>m\<close> and \<open>n\<close>. The
conjunctions don't seem provable separately.
\<close>
@@ -130,7 +130,7 @@
\<close>
lemma gcd_mult_distrib2: "k * gcd m n = gcd (k * m) (k * n)"
- -- \<open>@{cite \<open>page 27\<close> davenport92}\<close>
+ \<comment> \<open>@{cite \<open>page 27\<close> davenport92}\<close>
apply (induct m n rule: gcd_induct)
apply simp
apply (case_tac "k = 0")
@@ -697,7 +697,7 @@
done
lemmas zgcd_ac = zgcd_assoc zgcd_commute zgcd_left_commute
- -- \<open>addition is an AC-operator\<close>
+ \<comment> \<open>addition is an AC-operator\<close>
lemma zgcd_zmult_distrib2: "0 \<le> k ==> k * zgcd m n = zgcd (k * m) (k * n)"
by (simp del: minus_mult_right [symmetric]
--- a/src/HOL/Old_Number_Theory/WilsonBij.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Old_Number_Theory/WilsonBij.thy Thu May 26 17:51:22 2016 +0200
@@ -47,7 +47,7 @@
lemma inv_not_0:
"zprime p ==> 1 < a ==> a < p - 1 ==> inv p a \<noteq> 0"
- -- \<open>same as @{text WilsonRuss}\<close>
+ \<comment> \<open>same as \<open>WilsonRuss\<close>\<close>
apply safe
apply (cut_tac a = a and p = p in inv_is_inv)
apply (unfold zcong_def)
@@ -56,7 +56,7 @@
lemma inv_not_1:
"zprime p ==> 1 < a ==> a < p - 1 ==> inv p a \<noteq> 1"
- -- \<open>same as @{text WilsonRuss}\<close>
+ \<comment> \<open>same as \<open>WilsonRuss\<close>\<close>
apply safe
apply (cut_tac a = a and p = p in inv_is_inv)
prefer 4
@@ -67,7 +67,7 @@
done
lemma aux: "[a * (p - 1) = 1] (mod p) = [a = p - 1] (mod p)"
- -- \<open>same as @{text WilsonRuss}\<close>
+ \<comment> \<open>same as \<open>WilsonRuss\<close>\<close>
apply (unfold zcong_def)
apply (simp add: diff_diff_eq diff_diff_eq2 right_diff_distrib)
apply (rule_tac s = "p dvd -((a + 1) + (p * -a))" in trans)
@@ -81,7 +81,7 @@
lemma inv_not_p_minus_1:
"zprime p ==> 1 < a ==> a < p - 1 ==> inv p a \<noteq> p - 1"
- -- \<open>same as @{text WilsonRuss}\<close>
+ \<comment> \<open>same as \<open>WilsonRuss\<close>\<close>
apply safe
apply (cut_tac a = a and p = p in inv_is_inv)
apply auto
@@ -93,7 +93,7 @@
text \<open>
Below is slightly different as we don't expand @{term [source] inv}
- but use ``@{text correct}'' theorems.
+ but use ``\<open>correct\<close>'' theorems.
\<close>
lemma inv_g_1: "zprime p ==> 1 < a ==> a < p - 1 ==> 1 < inv p a"
@@ -111,7 +111,7 @@
lemma inv_less_p_minus_1:
"zprime p ==> 1 < a ==> a < p - 1 ==> inv p a < p - 1"
- -- \<open>ditto\<close>
+ \<comment> \<open>ditto\<close>
apply (subst order_less_le)
apply (simp add: inv_not_p_minus_1 inv_less)
done
--- a/src/HOL/Predicate_Compile_Examples/Code_Prolog_Examples.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Predicate_Compile_Examples/Code_Prolog_Examples.thy Thu May 26 17:51:22 2016 +0200
@@ -2,7 +2,7 @@
imports "~~/src/HOL/Library/Code_Prolog"
begin
-section {* Example append *}
+section \<open>Example append\<close>
inductive append
@@ -10,13 +10,13 @@
"append [] ys ys"
| "append xs ys zs ==> append (x # xs) ys (x # zs)"
-setup {* Code_Prolog.map_code_options (K
+setup \<open>Code_Prolog.map_code_options (K
{ensure_groundness = false,
limit_globally = NONE,
limited_types = [],
limited_predicates = [],
replacing = [],
- manual_reorder = []}) *}
+ manual_reorder = []})\<close>
values_prolog "{(x, y, z). append x y z}"
@@ -24,73 +24,73 @@
values_prolog 3 "{(x, y, z). append x y z}"
-setup {* Code_Prolog.map_code_options (K
+setup \<open>Code_Prolog.map_code_options (K
{ensure_groundness = false,
limit_globally = NONE,
limited_types = [],
limited_predicates = [],
replacing = [],
- manual_reorder = []}) *}
+ manual_reorder = []})\<close>
values_prolog "{(x, y, z). append x y z}"
-setup {* Code_Prolog.map_code_options (K
+setup \<open>Code_Prolog.map_code_options (K
{ensure_groundness = false,
limit_globally = NONE,
limited_types = [],
limited_predicates = [],
replacing = [],
- manual_reorder = []}) *}
+ manual_reorder = []})\<close>
-section {* Example queens *}
+section \<open>Example queens\<close>
inductive nodiag :: "int => int => int list => bool"
where
"nodiag B D []"
| "D \<noteq> N - B ==> D \<noteq> B - N ==> Da = D + 1 ==> nodiag B Da L ==> nodiag B D (N # L)"
-text {*
+text \<open>
qdelete(A, [A|L], L).
qdelete(X, [A|Z], [A|R]) :-
qdelete(X, Z, R).
-*}
+\<close>
inductive qdelete :: "int => int list => int list => bool"
where
"qdelete A (A # L) L"
| "qdelete X Z R ==> qdelete X (A # Z) (A # R)"
-text {*
+text \<open>
qperm([], []).
qperm([X|Y], K) :-
qdelete(U, [X|Y], Z),
K = [U|V],
qperm(Z, V).
-*}
+\<close>
inductive qperm :: "int list => int list => bool"
where
"qperm [] []"
| "qdelete U (X # Y) Z ==> qperm Z V ==> qperm (X # Y) (U # V)"
-text {*
+text \<open>
safe([]).
safe([N|L]) :-
nodiag(N, 1, L),
safe(L).
-*}
+\<close>
inductive safe :: "int list => bool"
where
"safe []"
| "nodiag N 1 L ==> safe L ==> safe (N # L)"
-text {*
+text \<open>
queen(Data, Out) :-
qperm(Data, Out),
safe(Out)
-*}
+\<close>
inductive queen :: "int list => int list => bool"
where
@@ -103,14 +103,14 @@
values_prolog 10 "{ys. queen_9 ys}"
-section {* Example symbolic derivation *}
+section \<open>Example symbolic derivation\<close>
hide_const Pow
datatype expr = Log expr | Mult expr expr | Div expr expr | x | Num int | Plus expr expr
| Minus expr expr | Uminus expr | Pow expr int | Exp expr
-text {*
+text \<open>
d(U + V, X, DU + DV) :-
cut,
@@ -145,7 +145,7 @@
cut.
d(num(_), _, num(0)).
-*}
+\<close>
inductive d :: "expr => expr => expr => bool"
where
@@ -160,7 +160,7 @@
| "d x X (Num 1)"
| "d (Num n) X (Num 0)"
-text {*
+text \<open>
ops8(E) :-
d((x + num(1)) * ((^(x, 2) + num(2)) * (^(x, 3) + num(3))), x, E).
@@ -172,7 +172,7 @@
times10(E) :-
d(((((((((x * x) * x) * x) * x) * x) * x) * x) * x) * x, x, E)
-*}
+\<close>
inductive ops8 :: "expr => bool"
where
@@ -195,7 +195,7 @@
values_prolog "{e. log10 e}"
values_prolog "{e. times10 e}"
-section {* Example negation *}
+section \<open>Example negation\<close>
datatype abc = A | B | C
@@ -203,13 +203,13 @@
where
"y \<noteq> B \<Longrightarrow> notB y"
-setup {* Code_Prolog.map_code_options (K
+setup \<open>Code_Prolog.map_code_options (K
{ensure_groundness = true,
limit_globally = NONE,
limited_types = [],
limited_predicates = [],
replacing = [],
- manual_reorder = []}) *}
+ manual_reorder = []})\<close>
values_prolog 2 "{y. notB y}"
@@ -219,7 +219,7 @@
values_prolog 5 "{y. notAB y}"
-section {* Example prolog conform variable names *}
+section \<open>Example prolog conform variable names\<close>
inductive equals :: "abc => abc => bool"
where
--- a/src/HOL/Predicate_Compile_Examples/Context_Free_Grammar_Example.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Predicate_Compile_Examples/Context_Free_Grammar_Example.thy Thu May 26 17:51:22 2016 +0200
@@ -5,7 +5,7 @@
declare mem_def[code_pred_inline]
*)
-subsection {* Alternative rules for length *}
+subsection \<open>Alternative rules for length\<close>
definition size_list :: "'a list => nat"
where "size_list = size"
@@ -19,10 +19,10 @@
declare size_list_def[symmetric, code_pred_inline]
-setup {*
+setup \<open>
Context.theory_map
(Quickcheck.add_tester ("prolog", (Code_Prolog.active, Code_Prolog.test_goals)))
-*}
+\<close>
datatype alphabet = a | b
@@ -59,13 +59,13 @@
declare filter_b_def[symmetric, code_pred_inline]
-setup {* Code_Prolog.map_code_options (K
+setup \<open>Code_Prolog.map_code_options (K
{ensure_groundness = true,
limit_globally = NONE,
limited_types = [],
limited_predicates = [(["s1p", "a1p", "b1p"], 2)],
replacing = [(("s1p", "limited_s1p"), "quickcheck")],
- manual_reorder = [(("quickcheck", 1), [0,2,1,4,3,5])]}) *}
+ manual_reorder = [(("quickcheck", 1), [0,2,1,4,3,5])]})\<close>
theorem S\<^sub>1_sound:
@@ -83,13 +83,13 @@
| "\<lbrakk>v \<in> B\<^sub>2; v \<in> B\<^sub>2\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^sub>2"
-setup {* Code_Prolog.map_code_options (K
+setup \<open>Code_Prolog.map_code_options (K
{ensure_groundness = true,
limit_globally = NONE,
limited_types = [],
limited_predicates = [(["s2p", "a2p", "b2p"], 3)],
replacing = [(("s2p", "limited_s2p"), "quickcheck")],
- manual_reorder = [(("quickcheck", 1), [0,2,1,4,3,5])]}) *}
+ manual_reorder = [(("quickcheck", 1), [0,2,1,4,3,5])]})\<close>
theorem S\<^sub>2_sound:
@@ -106,13 +106,13 @@
| "\<lbrakk>v \<in> B\<^sub>3; w \<in> B\<^sub>3\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^sub>3"
-setup {* Code_Prolog.map_code_options (K
+setup \<open>Code_Prolog.map_code_options (K
{ensure_groundness = true,
limit_globally = NONE,
limited_types = [],
limited_predicates = [(["s3p", "a3p", "b3p"], 6)],
replacing = [(("s3p", "limited_s3p"), "quickcheck")],
- manual_reorder = [(("quickcheck", 1), [0,2,1,4,3,5])]}) *}
+ manual_reorder = [(("quickcheck", 1), [0,2,1,4,3,5])]})\<close>
lemma S\<^sub>3_sound:
"S\<^sub>3p w \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"
@@ -148,13 +148,13 @@
| "\<lbrakk>v \<in> B\<^sub>4; w \<in> B\<^sub>4\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^sub>4"
-setup {* Code_Prolog.map_code_options (K
+setup \<open>Code_Prolog.map_code_options (K
{ensure_groundness = true,
limit_globally = NONE,
limited_types = [],
limited_predicates = [(["s4p", "a4p", "b4p"], 6)],
replacing = [(("s4p", "limited_s4p"), "quickcheck")],
- manual_reorder = [(("quickcheck", 1), [0,2,1,4,3,5])]}) *}
+ manual_reorder = [(("quickcheck", 1), [0,2,1,4,3,5])]})\<close>
theorem S\<^sub>4_sound:
--- a/src/HOL/Predicate_Compile_Examples/Examples.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Predicate_Compile_Examples/Examples.thy Thu May 26 17:51:22 2016 +0200
@@ -4,11 +4,11 @@
declare [[values_timeout = 480.0]]
-section {* Formal Languages *}
+section \<open>Formal Languages\<close>
-subsection {* General Context Free Grammars *}
+subsection \<open>General Context Free Grammars\<close>
-text {* a contribution by Aditi Barthwal *}
+text \<open>a contribution by Aditi Barthwal\<close>
datatype ('nts,'ts) symbol = NTS 'nts
| TS 'ts
@@ -73,7 +73,7 @@
values "{rhs. test2 rhs}"
-subsection {* Some concrete Context Free Grammars *}
+subsection \<open>Some concrete Context Free Grammars\<close>
datatype alphabet = a | b
@@ -133,9 +133,9 @@
hide_const a b
-section {* Semantics of programming languages *}
+section \<open>Semantics of programming languages\<close>
-subsection {* IMP *}
+subsection \<open>IMP\<close>
type_synonym var = nat
type_synonym state = "int list"
@@ -162,7 +162,7 @@
(While (%s. s!0 > 0) (Seq (Ass 0 (%s. s!0 - 1)) (Ass 1 (%s. s!1 + 1))))
[3,5] t}"
-subsection {* Lambda *}
+subsection \<open>Lambda\<close>
datatype type =
Atom nat
@@ -229,13 +229,13 @@
values [random_dseq 1,1,5] 10 "{(\<Gamma>, t, T). \<Gamma> \<turnstile> t : T}"
-subsection {* A minimal example of yet another semantics *}
+subsection \<open>A minimal example of yet another semantics\<close>
-text {* thanks to Elke Salecker *}
+text \<open>thanks to Elke Salecker\<close>
type_synonym vname = nat
type_synonym vvalue = int
-type_synonym var_assign = "vname \<Rightarrow> vvalue" --"variable assignment"
+type_synonym var_assign = "vname \<Rightarrow> vvalue" \<comment>"variable assignment"
datatype ir_expr =
IrConst vvalue
@@ -262,9 +262,9 @@
values "{val. eval_var (Add (IrConst 1) (IrConst 2)) (| Env = (\<lambda>x. 0)|) val}"
-subsection {* Another semantics *}
+subsection \<open>Another semantics\<close>
-type_synonym name = nat --"For simplicity in examples"
+type_synonym name = nat \<comment>"For simplicity in examples"
type_synonym state' = "name \<Rightarrow> nat"
datatype aexp = N nat | V name | Plus aexp aexp
@@ -321,10 +321,10 @@
)))))))))))))))))))))))))))))))))))))))))]}"] "{list s 2|s. (SKIP, nth [42, 43]) \<Rightarrow> s}"
-subsection {* CCS *}
+subsection \<open>CCS\<close>
-text{* This example formalizes finite CCS processes without communication or
-recursion. For simplicity, labels are natural numbers. *}
+text\<open>This example formalizes finite CCS processes without communication or
+recursion. For simplicity, labels are natural numbers.\<close>
datatype proc = nil | pre nat proc | or proc proc | par proc proc
--- a/src/HOL/Predicate_Compile_Examples/Hotel_Example.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Predicate_Compile_Examples/Hotel_Example.thy Thu May 26 17:51:22 2016 +0200
@@ -81,7 +81,7 @@
no_Check_in (s\<^sub>3 @ s\<^sub>2) r \<and> isin (s\<^sub>2 @ [Check_in g r c] @ s\<^sub>1) r = {})"
-section {* Some setup *}
+section \<open>Some setup\<close>
lemma issued_nil: "issued [] = {Key0}"
by (auto simp add: initk_def)
--- a/src/HOL/Predicate_Compile_Examples/Hotel_Example_Prolog.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Predicate_Compile_Examples/Hotel_Example_Prolog.thy Thu May 26 17:51:22 2016 +0200
@@ -71,58 +71,58 @@
(\<not> (\<exists> g'. isinp (s\<^sub>2 @ [Check_in g r c] @ s\<^sub>1) r g')))"
unfolding feels_safe_def isinp_def by auto
-setup {* Code_Prolog.map_code_options (K
+setup \<open>Code_Prolog.map_code_options (K
{ensure_groundness = true,
limit_globally = NONE,
limited_types = [],
limited_predicates = [],
replacing = [],
- manual_reorder = []}) *}
+ manual_reorder = []})\<close>
values_prolog 40 "{s. hotel s}"
-setup {*
+setup \<open>
Context.theory_map
(Quickcheck.add_tester ("prolog", (Code_Prolog.active, Code_Prolog.test_goals)))
-*}
+\<close>
lemma "\<lbrakk> hotel s; isinp s r g \<rbrakk> \<Longrightarrow> owns s r = Some g"
quickcheck[tester = prolog, iterations = 1, expect = counterexample]
oops
-section {* Manual setup to find the counterexample *}
+section \<open>Manual setup to find the counterexample\<close>
-setup {* Code_Prolog.map_code_options (K
+setup \<open>Code_Prolog.map_code_options (K
{ensure_groundness = true,
limit_globally = NONE,
limited_types = [],
limited_predicates = [(["hotel"], 4)],
replacing = [(("hotel", "limited_hotel"), "quickcheck")],
- manual_reorder = []}) *}
+ manual_reorder = []})\<close>
lemma
"hotel s ==> feels_safe s r ==> isinp s r g ==> owns s r = Some g"
quickcheck[tester = prolog, iterations = 1, expect = no_counterexample]
oops
-setup {* Code_Prolog.map_code_options (K
+setup \<open>Code_Prolog.map_code_options (K
{ensure_groundness = true,
limit_globally = NONE,
limited_types = [],
limited_predicates = [(["hotel"], 5)],
replacing = [(("hotel", "limited_hotel"), "quickcheck")],
- manual_reorder = []}) *}
+ manual_reorder = []})\<close>
lemma
"hotel s ==> feels_safe s r ==> isinp s r g ==> owns s r = Some g"
quickcheck[tester = prolog, iterations = 1, expect = counterexample]
oops
-section {* Using a global limit for limiting the execution *}
+section \<open>Using a global limit for limiting the execution\<close>
-text {* A global depth limit of 10 does not suffice to find the counterexample. *}
+text \<open>A global depth limit of 10 does not suffice to find the counterexample.\<close>
-setup {*
+setup \<open>
Code_Prolog.map_code_options (K
{ensure_groundness = true,
limit_globally = SOME 10,
@@ -130,16 +130,16 @@
limited_predicates = [],
replacing = [],
manual_reorder = []})
-*}
+\<close>
lemma
"hotel s ==> feels_safe s r ==> isinp s r g ==> owns s r = Some g"
quickcheck[tester = prolog, iterations = 1, expect = no_counterexample]
oops
-text {* But a global depth limit of 11 does. *}
+text \<open>But a global depth limit of 11 does.\<close>
-setup {*
+setup \<open>
Code_Prolog.map_code_options (K
{ensure_groundness = true,
limit_globally = SOME 11,
@@ -147,7 +147,7 @@
limited_predicates = [],
replacing = [],
manual_reorder = []})
-*}
+\<close>
lemma
"hotel s ==> feels_safe s r ==> isinp s r g ==> owns s r = Some g"
--- a/src/HOL/Predicate_Compile_Examples/Hotel_Example_Small_Generator.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Predicate_Compile_Examples/Hotel_Example_Small_Generator.thy Thu May 26 17:51:22 2016 +0200
@@ -42,19 +42,19 @@
end
-ML {*
+ML \<open>
val small_15_active = Attrib.setup_config_bool @{binding quickcheck_small_14_active} (K false);
val small_14_active = Attrib.setup_config_bool @{binding quickcheck_small_15_active} (K false);
-*}
+\<close>
-setup {*
+setup \<open>
Context.theory_map (Quickcheck.add_tester ("small_generators_depth_14",
(small_14_active, Predicate_Compile_Quickcheck.test_goals
(Predicate_Compile_Aux.Pos_Generator_DSeq, true, true, 14))))
#> Context.theory_map (Quickcheck.add_tester ("small_generators_depth_15",
(small_15_active, Predicate_Compile_Quickcheck.test_goals
(Predicate_Compile_Aux.Pos_Generator_DSeq, true, true, 15))))
-*}
+\<close>
lemma
"hotel s ==> feels_safe s r ==> g \<in> isin s r ==> owns s r = Some g"
--- a/src/HOL/Predicate_Compile_Examples/IMP_1.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Predicate_Compile_Examples/IMP_1.thy Thu May 26 17:51:22 2016 +0200
@@ -2,11 +2,11 @@
imports "~~/src/HOL/Library/Predicate_Compile_Quickcheck"
begin
-subsection {* IMP *}
+subsection \<open>IMP\<close>
-text {*
+text \<open>
In this example, the state is one boolean variable and the commands are Skip, Ass, Seq, and IF.
-*}
+\<close>
type_synonym var = unit
type_synonym state = bool
--- a/src/HOL/Predicate_Compile_Examples/IMP_2.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Predicate_Compile_Examples/IMP_2.thy Thu May 26 17:51:22 2016 +0200
@@ -2,11 +2,11 @@
imports "~~/src/HOL/Library/Predicate_Compile_Quickcheck"
begin
-subsection {* IMP *}
+subsection \<open>IMP\<close>
-text {*
+text \<open>
In this example, the state is one boolean variable and the commands are Skip, Ass, Seq, IF and While.
-*}
+\<close>
type_synonym var = unit
type_synonym state = bool
--- a/src/HOL/Predicate_Compile_Examples/IMP_3.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Predicate_Compile_Examples/IMP_3.thy Thu May 26 17:51:22 2016 +0200
@@ -2,11 +2,11 @@
imports "~~/src/HOL/Library/Predicate_Compile_Quickcheck"
begin
-subsection {* IMP *}
+subsection \<open>IMP\<close>
-text {*
+text \<open>
In this example, the state is one integer variable and the commands are Skip, Ass, Seq, IF, and While.
-*}
+\<close>
type_synonym var = unit
type_synonym state = int
--- a/src/HOL/Predicate_Compile_Examples/IMP_4.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Predicate_Compile_Examples/IMP_4.thy Thu May 26 17:51:22 2016 +0200
@@ -2,11 +2,11 @@
imports "~~/src/HOL/Library/Predicate_Compile_Quickcheck"
begin
-subsection {* IMP *}
+subsection \<open>IMP\<close>
-text {*
+text \<open>
In this example, the state is a list of integers and the commands are Skip, Ass, Seq, IF and While.
-*}
+\<close>
type_synonym var = nat
type_synonym state = "int list"
--- a/src/HOL/Predicate_Compile_Examples/Lambda_Example.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Predicate_Compile_Examples/Lambda_Example.thy Thu May 26 17:51:22 2016 +0200
@@ -2,7 +2,7 @@
imports "~~/src/HOL/Library/Code_Prolog"
begin
-subsection {* Lambda *}
+subsection \<open>Lambda\<close>
datatype type =
Atom nat
@@ -56,7 +56,7 @@
| appR [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> u \<degree> s \<rightarrow>\<^sub>\<beta> u \<degree> t"
| abs [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> Abs T s \<rightarrow>\<^sub>\<beta> Abs T t"
-subsection {* Inductive definitions for ordering on naturals *}
+subsection \<open>Inductive definitions for ordering on naturals\<close>
inductive less_nat
where
@@ -77,28 +77,28 @@
lemma [code_pred_inline]: "(x::nat) + 1 = Suc x"
by simp
-section {* Manual setup to find counterexample *}
+section \<open>Manual setup to find counterexample\<close>
-setup {*
+setup \<open>
Context.theory_map
(Quickcheck.add_tester ("prolog", (Code_Prolog.active, Code_Prolog.test_goals)))
-*}
+\<close>
-setup {* Code_Prolog.map_code_options (K
+setup \<open>Code_Prolog.map_code_options (K
{ ensure_groundness = true,
limit_globally = NONE,
limited_types = [(@{typ nat}, 1), (@{typ "type"}, 1), (@{typ dB}, 1), (@{typ "type list"}, 1)],
limited_predicates = [(["typing"], 2), (["nthel1"], 2)],
replacing = [(("typing", "limited_typing"), "quickcheck"),
(("nthel1", "limited_nthel1"), "lim_typing")],
- manual_reorder = []}) *}
+ manual_reorder = []})\<close>
lemma
"\<Gamma> \<turnstile> t : U \<Longrightarrow> t \<rightarrow>\<^sub>\<beta> t' \<Longrightarrow> \<Gamma> \<turnstile> t' : U"
quickcheck[tester = prolog, iterations = 1, expect = counterexample]
oops
-text {* Verifying that the found counterexample really is one by means of a proof *}
+text \<open>Verifying that the found counterexample really is one by means of a proof\<close>
lemma
assumes
--- a/src/HOL/Predicate_Compile_Examples/List_Examples.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Predicate_Compile_Examples/List_Examples.thy Thu May 26 17:51:22 2016 +0200
@@ -5,12 +5,12 @@
"~~/src/HOL/Library/Code_Prolog"
begin
-setup {*
+setup \<open>
Context.theory_map
(Quickcheck.add_tester ("prolog", (Code_Prolog.active, Code_Prolog.test_goals)))
-*}
+\<close>
-setup {* Code_Prolog.map_code_options (K
+setup \<open>Code_Prolog.map_code_options (K
{ensure_groundness = true,
limit_globally = NONE,
limited_types = [(@{typ nat}, 2), (@{typ "nat list"}, 4)],
@@ -19,7 +19,7 @@
[(("appendP", "limited_appendP"), "quickcheck"),
(("revP", "limited_revP"), "quickcheck"),
(("appendP", "limited_appendP"), "lim_revP")],
- manual_reorder = []}) *}
+ manual_reorder = []})\<close>
lemma "(xs :: nat list) = ys @ ys --> rev xs = xs"
quickcheck[tester = random, iterations = 10000]
--- a/src/HOL/Predicate_Compile_Examples/Predicate_Compile_Quickcheck_Examples.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Predicate_Compile_Examples/Predicate_Compile_Quickcheck_Examples.thy Thu May 26 17:51:22 2016 +0200
@@ -38,7 +38,7 @@
oops
*)
-section {* Equivalences *}
+section \<open>Equivalences\<close>
inductive is_ten :: "nat => bool"
where
@@ -53,7 +53,7 @@
quickcheck[tester = smart_exhaustive, iterations = 1, size = 1, expect = counterexample]
oops
-section {* Context Free Grammar *}
+section \<open>Context Free Grammar\<close>
datatype alphabet = a | b
@@ -169,14 +169,14 @@
hide_const a b
-subsection {* Lexicographic order *}
+subsection \<open>Lexicographic order\<close>
(* TODO *)
(*
lemma
"(u, v) : lexord r ==> (x @ u, y @ v) : lexord r"
oops
*)
-subsection {* IMP *}
+subsection \<open>IMP\<close>
type_synonym var = nat
type_synonym state = "int list"
@@ -206,7 +206,7 @@
quickcheck[tester = smart_exhaustive, size=2, iterations=20, expect = counterexample]
oops
-subsection {* Lambda *}
+subsection \<open>Lambda\<close>
datatype type =
Atom nat
@@ -261,7 +261,7 @@
quickcheck[tester = smart_exhaustive, size = 7, iterations = 10]
oops
-subsection {* JAD *}
+subsection \<open>JAD\<close>
definition matrix :: "('a :: semiring_0) list list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" where
"matrix M rs cs \<longleftrightarrow> (\<forall> row \<in> set M. length row = cs) \<and> length M = rs"
@@ -294,10 +294,10 @@
definition mv :: "('a :: semiring_0) list list \<Rightarrow> 'a list \<Rightarrow> 'a list"
where [simp]: "mv M v = map (scalar_product v) M"
-text {*
+text \<open>
This defines the matrix vector multiplication. To work properly @{term
"matrix M m n \<and> length v = n"} must hold.
-*}
+\<close>
subsection "Compressed matrix"
--- a/src/HOL/Predicate_Compile_Examples/Predicate_Compile_Tests.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Predicate_Compile_Examples/Predicate_Compile_Tests.thy Thu May 26 17:51:22 2016 +0200
@@ -4,7 +4,7 @@
declare [[values_timeout = 480.0]]
-subsection {* Basic predicates *}
+subsection \<open>Basic predicates\<close>
inductive False' :: "bool"
@@ -89,16 +89,16 @@
code_pred (expected_modes: o => o => bool, i => o => bool, o => i => bool, i => i => bool) Fact .
-text {*
+text \<open>
The following example was derived from an predicate in the CakeML formalisation provided by Lars Hupel.
-*}
+\<close>
inductive predicate_where_argument_is_condition :: "bool \<Rightarrow> bool"
where
"ck \<Longrightarrow> predicate_where_argument_is_condition ck"
code_pred predicate_where_argument_is_condition .
-text {* Other similar examples of this kind: *}
+text \<open>Other similar examples of this kind:\<close>
inductive predicate_where_argument_is_in_equation :: "bool \<Rightarrow> bool"
where
@@ -138,10 +138,10 @@
thm zerozero.dseq_equation
thm zerozero.random_dseq_equation
-text {* We expect the user to expand the tuples in the values command.
-The following values command is not supported. *}
+text \<open>We expect the user to expand the tuples in the values command.
+The following values command is not supported.\<close>
(*values "{x. zerozero x}" *)
-text {* Instead, the user must type *}
+text \<open>Instead, the user must type\<close>
values "{(x, y). zerozero (x, y)}"
values [expected "{}" dseq 0] "{(x, y). zerozero (x, y)}"
@@ -174,7 +174,7 @@
values [expected "{(0::nat, 7::natural)}"] "{(a, c). JamesBond a 0 c}"
-subsection {* Alternative Rules *}
+subsection \<open>Alternative Rules\<close>
datatype char = C | D | E | F | G | H
@@ -233,7 +233,7 @@
"is_F_or_G x ==> is_FGH x"
| "is_FGH H"
-text {* Compilation of is_FGH requires elimination rule for is_F_or_G *}
+text \<open>Compilation of is_FGH requires elimination rule for is_F_or_G\<close>
code_pred (expected_modes: o => bool, i => bool) is_FGH
proof -
@@ -254,7 +254,7 @@
qed
qed
-subsection {* Named alternative rules *}
+subsection \<open>Named alternative rules\<close>
inductive appending
where
@@ -267,8 +267,8 @@
lemma appending_alt_cons: "xs' = x # xs \<Longrightarrow> appending xs ys zs \<Longrightarrow> zs' = x # zs \<Longrightarrow> appending xs' ys zs'"
by (auto intro: appending.intros)
-text {* With code_pred_intro, we can give fact names to the alternative rules,
- which are used for the code_pred command. *}
+text \<open>With code_pred_intro, we can give fact names to the alternative rules,
+ which are used for the code_pred command.\<close>
declare appending_alt_nil[code_pred_intro alt_nil] appending_alt_cons[code_pred_intro alt_cons]
@@ -321,7 +321,7 @@
qed
qed
-subsection {* Preprocessor Inlining *}
+subsection \<open>Preprocessor Inlining\<close>
definition "equals == (op =)"
@@ -346,11 +346,11 @@
definition "zerozero'' x == zerozero' x"
-text {* if preprocessing fails, zerozero'' will not have all modes. *}
+text \<open>if preprocessing fails, zerozero'' will not have all modes.\<close>
code_pred (expected_modes: i * i => bool, i * o => bool, o * i => bool, o => bool) [inductify] zerozero'' .
-subsection {* Sets *}
+subsection \<open>Sets\<close>
(*
inductive_set EmptySet :: "'a set"
@@ -361,7 +361,7 @@
code_pred (expected_modes: o => bool, i => bool) [inductify] EmptySet' .
*)
-subsection {* Numerals *}
+subsection \<open>Numerals\<close>
definition
"one_or_two = (%x. x = Suc 0 \<or> ( x = Suc (Suc 0)))"
@@ -392,7 +392,7 @@
values "{x. one_or_two'' x}"
-subsection {* even predicate *}
+subsection \<open>even predicate\<close>
inductive even :: "nat \<Rightarrow> bool" and odd :: "nat \<Rightarrow> bool" where
"even 0"
@@ -456,7 +456,7 @@
code_pred (expected_modes: i => bool) is_even .
-subsection {* append predicate *}
+subsection \<open>append predicate\<close>
inductive append :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool" where
"append [] xs xs"
@@ -493,7 +493,7 @@
value "Predicate.the (slice ([]::int list))"
-text {* tricky case with alternative rules *}
+text \<open>tricky case with alternative rules\<close>
inductive append2
where
@@ -565,7 +565,7 @@
i * o * i => bool, i * i * i => bool) tupled_append''' .
thm tupled_append'''.equation
-subsection {* map_ofP predicate *}
+subsection \<open>map_ofP predicate\<close>
inductive map_ofP :: "('a \<times> 'b) list \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool"
where
@@ -575,7 +575,7 @@
code_pred (expected_modes: i => o => o => bool, i => i => o => bool, i => o => i => bool, i => i => i => bool) map_ofP .
thm map_ofP.equation
-subsection {* filter predicate *}
+subsection \<open>filter predicate\<close>
inductive filter1
for P
@@ -626,7 +626,7 @@
(*code_pred [depth_limited] filter4 .*)
(*code_pred [random] filter4 .*)
*)
-subsection {* reverse predicate *}
+subsection \<open>reverse predicate\<close>
inductive rev where
"rev [] []"
@@ -645,7 +645,7 @@
code_pred (expected_modes: i * o => bool, o * i => bool, i * i => bool) tupled_rev .
thm tupled_rev.equation
-subsection {* partition predicate *}
+subsection \<open>partition predicate\<close>
inductive partition :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool"
for f where
@@ -680,9 +680,9 @@
"r a b \<Longrightarrow> tranclp r b c \<Longrightarrow> tranclp r a c"
by auto
-subsection {* transitive predicate *}
+subsection \<open>transitive predicate\<close>
-text {* Also look at the tabled transitive closure in the Library *}
+text \<open>Also look at the tabled transitive closure in the Library\<close>
code_pred (modes: (i => o => bool) => i => i => bool, (i => o => bool) => i => o => bool as forwards_trancl,
(o => i => bool) => i => i => bool, (o => i => bool) => o => i => bool as backwards_trancl, (o => o => bool) => i => i => bool, (o => o => bool) => i => o => bool,
@@ -735,8 +735,8 @@
values "{m. succ 0 m}"
values "{m. succ m 0}"
-text {* values command needs mode annotation of the parameter succ
-to disambiguate which mode is to be chosen. *}
+text \<open>values command needs mode annotation of the parameter succ
+to disambiguate which mode is to be chosen.\<close>
values [mode: i => o => bool] 20 "{n. tranclp succ 10 n}"
values [mode: o => i => bool] 10 "{n. tranclp succ n 10}"
@@ -763,7 +763,7 @@
value "not_reachable_in_example_graph 0 3"
value "not_reachable_in_example_graph 4 8"
value "not_reachable_in_example_graph 5 6"
-text {* rtrancl compilation is strange! *}
+text \<open>rtrancl compilation is strange!\<close>
(*
value "not_reachable_in_example_graph 0 4"
value "not_reachable_in_example_graph 1 6"
@@ -808,7 +808,7 @@
(*values [depth_limit = 4] "{x. not_reachable_in_example_graph' 0 4}"*) (* fails with undefined *)
(*values [depth_limit = 20] "{x. not_reachable_in_example_graph' 0 4}"*) (* fails with undefined *)
-subsection {* Free function variable *}
+subsection \<open>Free function variable\<close>
inductive FF :: "nat => nat => bool"
where
@@ -816,7 +816,7 @@
code_pred FF .
-subsection {* IMP *}
+subsection \<open>IMP\<close>
type_synonym var = nat
type_synonym state = "int list"
@@ -841,10 +841,10 @@
values "{s. tupled_exec (While (%s. s!0 > 0) (Seq (Ass 0 (%s. s!0 - 1)) (Ass 1 (%s. s!1 + 1))), [3, 5], s)}"
-subsection {* CCS *}
+subsection \<open>CCS\<close>
-text{* This example formalizes finite CCS processes without communication or
-recursion. For simplicity, labels are natural numbers. *}
+text\<open>This example formalizes finite CCS processes without communication or
+recursion. For simplicity, labels are natural numbers.\<close>
datatype proc = nil | pre nat proc | or proc proc | par proc proc
@@ -859,7 +859,7 @@
code_pred tupled_step .
thm tupled_step.equation
-subsection {* divmod *}
+subsection \<open>divmod\<close>
inductive divmod_rel :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" where
"k < l \<Longrightarrow> divmod_rel k l 0 k"
@@ -869,9 +869,9 @@
thm divmod_rel.equation
value "Predicate.the (divmod_rel_i_i_o_o 1705 42)"
-subsection {* Transforming predicate logic into logic programs *}
+subsection \<open>Transforming predicate logic into logic programs\<close>
-subsection {* Transforming functions into logic programs *}
+subsection \<open>Transforming functions into logic programs\<close>
definition
"case_f xs ys = (case (xs @ ys) of [] => [] | (x # xs) => xs)"
@@ -886,7 +886,7 @@
code_pred [inductify] fold_map_idx .
-subsection {* Minimum *}
+subsection \<open>Minimum\<close>
definition Min
where "Min s r x \<equiv> s x \<and> (\<forall>y. r x y \<longrightarrow> x = y)"
@@ -894,9 +894,9 @@
code_pred [inductify] Min .
thm Min.equation
-subsection {* Lexicographic order *}
-text {* This example requires to handle the differences of sets and predicates in the predicate compiler,
-or to have a copy of all definitions on predicates due to the set-predicate distinction. *}
+subsection \<open>Lexicographic order\<close>
+text \<open>This example requires to handle the differences of sets and predicates in the predicate compiler,
+or to have a copy of all definitions on predicates due to the set-predicate distinction.\<close>
(*
declare lexord_def[code_pred_def]
@@ -1009,7 +1009,7 @@
code_pred [inductify] lists .
thm lists.equation
*)
-subsection {* AVL Tree *}
+subsection \<open>AVL Tree\<close>
datatype 'a tree = ET | MKT 'a "'a tree" "'a tree" nat
fun height :: "'a tree => nat" where
@@ -1050,7 +1050,7 @@
thm is_ord_aux.equation
thm is_ord.equation
*)
-subsection {* Definitions about Relations *}
+subsection \<open>Definitions about Relations\<close>
(*
code_pred (modes:
(i * i => bool) => i * i => bool,
@@ -1112,7 +1112,7 @@
code_pred [inductify] inv_image .
thm inv_image.equation
*)
-subsection {* Inverting list functions *}
+subsection \<open>Inverting list functions\<close>
code_pred [inductify, skip_proof] size_list' .
code_pred [new_random_dseq inductify] size_list' .
@@ -1203,12 +1203,12 @@
code_pred [inductify] filter .
code_pred [random_dseq inductify] filter .
-section {* Function predicate replacement *}
+section \<open>Function predicate replacement\<close>
-text {*
+text \<open>
If the mode analysis uses the functional mode, we
replace predicates that resulted from functions again by their functions.
-*}
+\<close>
inductive test_append
where
@@ -1217,13 +1217,13 @@
code_pred [inductify, skip_proof] test_append .
thm test_append.equation
-text {* If append is not turned into a predicate, then the mode
- o => o => i => bool could not be inferred. *}
+text \<open>If append is not turned into a predicate, then the mode
+ o => o => i => bool could not be inferred.\<close>
values 4 "{(xs::int list, ys). test_append xs ys [1, 2, 3, 4]}"
-text {* If appendP is not reverted back to a function, then mode i => i => o => bool
- fails after deleting the predicate equation. *}
+text \<open>If appendP is not reverted back to a function, then mode i => i => o => bool
+ fails after deleting the predicate equation.\<close>
declare appendP.equation[code del]
@@ -1231,11 +1231,11 @@
values "{xs::int list. test_append (replicate 1000 1) (replicate 1000 2) xs}"
values "{xs::int list. test_append (replicate 2000 1) (replicate 2000 2) xs}"
-text {* Redeclaring append.equation as code equation *}
+text \<open>Redeclaring append.equation as code equation\<close>
declare appendP.equation[code]
-subsection {* Function with tuples *}
+subsection \<open>Function with tuples\<close>
fun append'
where
@@ -1256,9 +1256,9 @@
values "{zs :: int list. test_append' [1,2,3] [4,5] zs}"
-section {* Arithmetic examples *}
+section \<open>Arithmetic examples\<close>
-subsection {* Examples on nat *}
+subsection \<open>Examples on nat\<close>
inductive plus_nat_test :: "nat => nat => nat => bool"
where
@@ -1304,7 +1304,7 @@
values [expected "{0, Suc 0, Suc (Suc 0), Suc (Suc (Suc 0))}"] "{x. minus_nat_test x 3 0}"
values [expected "{0::nat}"] "{x. minus_nat_test x 0 0}"
-subsection {* Examples on int *}
+subsection \<open>Examples on int\<close>
inductive plus_int_test :: "int => int => int => bool"
where
@@ -1334,7 +1334,7 @@
values [expected "{9::int}"] "{a. minus_int_test a 4 5}"
values [expected "{-1::int}"] "{b. minus_int_test 4 b 5}"
-subsection {* minus on bool *}
+subsection \<open>minus on bool\<close>
inductive All :: "nat => bool"
where
@@ -1349,7 +1349,7 @@
values "{x. test_minus_bool x}"
-subsection {* Functions *}
+subsection \<open>Functions\<close>
fun partial_hd :: "'a list => 'a option"
where
@@ -1364,7 +1364,7 @@
thm hd_predicate.equation
-subsection {* Locales *}
+subsection \<open>Locales\<close>
inductive hd_predicate2 :: "('a list => 'a option) => 'a list => 'a => bool"
where
@@ -1384,8 +1384,8 @@
end
-text {* The global introduction rules must be redeclared as introduction rules and then
- one could invoke code_pred. *}
+text \<open>The global introduction rules must be redeclared as introduction rules and then
+ one could invoke code_pred.\<close>
declare A.hd_predicate_in_locale.intros [code_pred_intro]
@@ -1394,12 +1394,12 @@
interpretation A partial_hd .
thm hd_predicate_in_locale.intros
-text {* A locally compliant solution with a trivial interpretation fails, because
-the predicate compiler has very strict assumptions about the terms and their structure. *}
+text \<open>A locally compliant solution with a trivial interpretation fails, because
+the predicate compiler has very strict assumptions about the terms and their structure.\<close>
(*code_pred hd_predicate_in_locale .*)
-section {* Integer example *}
+section \<open>Integer example\<close>
definition three :: int
where "three = 3"
@@ -1412,7 +1412,7 @@
thm is_three.equation
-section {* String.literal example *}
+section \<open>String.literal example\<close>
definition Error_1
where
@@ -1451,7 +1451,7 @@
thm is_error''.equation
-section {* Another function example *}
+section \<open>Another function example\<close>
consts f :: "'a \<Rightarrow> 'a"
@@ -1463,7 +1463,7 @@
thm fun_upd.equation
-section {* Examples for detecting switches *}
+section \<open>Examples for detecting switches\<close>
inductive detect_switches1 where
"detect_switches1 [] []"
@@ -1543,7 +1543,7 @@
thm detect_switches9.equation
-text {* The higher-order predicate r is in an output term *}
+text \<open>The higher-order predicate r is in an output term\<close>
datatype result = Result bool
@@ -1559,9 +1559,9 @@
thm test_relation_in_output_terms.equation
-text {*
+text \<open>
We want that the argument r is not treated as a higher-order relation, but simply as input.
-*}
+\<close>
inductive test_uninterpreted_relation :: "('a => bool) => 'a list => bool"
where
@@ -1583,14 +1583,14 @@
"list_ex r xs ==> test_uninterpreted_relation2 r xs"
| "list_ex' r xs ==> test_uninterpreted_relation2 r xs"
-text {* Proof procedure cannot handle this situation yet. *}
+text \<open>Proof procedure cannot handle this situation yet.\<close>
code_pred (modes: i => i => bool) [skip_proof] test_uninterpreted_relation2 .
thm test_uninterpreted_relation2.equation
-text {* Trivial predicate *}
+text \<open>Trivial predicate\<close>
inductive implies_itself :: "'a => bool"
where
@@ -1598,7 +1598,7 @@
code_pred implies_itself .
-text {* Case expressions *}
+text \<open>Case expressions\<close>
definition
"map_prods xs ys = (map (%((a, b), c). (a, b, c)) xs = ys)"
--- a/src/HOL/Predicate_Compile_Examples/Reg_Exp_Example.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Predicate_Compile_Examples/Reg_Exp_Example.thy Thu May 26 17:51:22 2016 +0200
@@ -4,10 +4,10 @@
"~~/src/HOL/Library/Code_Prolog"
begin
-text {* An example from the experiments from SmallCheck (@{url "http://www.cs.york.ac.uk/fp/smallcheck/"}) *}
-text {* The example (original in Haskell) was imported with Haskabelle and
+text \<open>An example from the experiments from SmallCheck (@{url "http://www.cs.york.ac.uk/fp/smallcheck/"})\<close>
+text \<open>The example (original in Haskell) was imported with Haskabelle and
manually slightly adapted.
-*}
+\<close>
datatype Nat = Zer
| Suc Nat
@@ -98,12 +98,12 @@
oops
-setup {*
+setup \<open>
Context.theory_map
(Quickcheck.add_tester ("prolog", (Code_Prolog.active, Code_Prolog.test_goals)))
-*}
+\<close>
-setup {* Code_Prolog.map_code_options (K
+setup \<open>Code_Prolog.map_code_options (K
{ensure_groundness = true,
limit_globally = NONE,
limited_types = [(@{typ Sym}, 0), (@{typ "Sym list"}, 2), (@{typ RE}, 6)],
@@ -114,10 +114,10 @@
(("subP", "limited_subP"), "repIntP"),
(("repIntPa", "limited_repIntPa"), "repIntP"),
(("accepts", "limited_accepts"), "quickcheck")],
- manual_reorder = []}) *}
+ manual_reorder = []})\<close>
-text {* Finding the counterexample still seems out of reach for the
- prolog-style generation. *}
+text \<open>Finding the counterexample still seems out of reach for the
+ prolog-style generation.\<close>
lemma "accepts (And (repInt n k p) (repInt n k q)) s --> accepts (repInt n k (And p q)) s"
quickcheck[exhaustive]
@@ -127,7 +127,7 @@
(*quickcheck[tester = prolog, iterations = 1, size = 1]*)
oops
-section {* Given a partial solution *}
+section \<open>Given a partial solution\<close>
lemma
(* assumes "n = Zer"
@@ -140,7 +140,7 @@
quickcheck[tester = prolog, iterations = 1, size = 1]
oops
-section {* Checking the counterexample *}
+section \<open>Checking the counterexample\<close>
definition a_sol
where
--- a/src/HOL/Predicate_Compile_Examples/Specialisation_Examples.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Predicate_Compile_Examples/Specialisation_Examples.thy Thu May 26 17:51:22 2016 +0200
@@ -4,7 +4,7 @@
declare [[values_timeout = 960.0]]
-section {* Specialisation Examples *}
+section \<open>Specialisation Examples\<close>
primrec nth_el'
where
@@ -15,18 +15,18 @@
"greater_than_index xs = (\<forall>i x. nth_el' xs i = Some x --> x > i)"
code_pred (expected_modes: i => bool) [inductify, skip_proof, specialise] greater_than_index .
-ML_val {* Core_Data.intros_of @{context} @{const_name specialised_nth_el'P} *}
+ML_val \<open>Core_Data.intros_of @{context} @{const_name specialised_nth_el'P}\<close>
thm greater_than_index.equation
values [expected "{()}"] "{x. greater_than_index [1,2,4,6]}"
values [expected "{}"] "{x. greater_than_index [0,2,3,2]}"
-subsection {* Common subterms *}
+subsection \<open>Common subterms\<close>
-text {* If a predicate is called with common subterms as arguments,
+text \<open>If a predicate is called with common subterms as arguments,
this predicate should be specialised.
-*}
+\<close>
definition max_nat :: "nat => nat => nat"
where "max_nat a b = (if a <= b then b else a)"
@@ -38,17 +38,17 @@
definition
"max_of_my_Suc x = max x (Suc x)"
-text {* In this example, max is specialised, hence the mode o => i => bool is possible *}
+text \<open>In this example, max is specialised, hence the mode o => i => bool is possible\<close>
code_pred (modes: o => i => bool) [inductify, specialise, skip_proof] max_of_my_Suc .
thm max_of_my_SucP.equation
-ML_val {* Core_Data.intros_of @{context} @{const_name specialised_max_natP} *}
+ML_val \<open>Core_Data.intros_of @{context} @{const_name specialised_max_natP}\<close>
values "{x. max_of_my_SucP x 6}"
-subsection {* Sorts *}
+subsection \<open>Sorts\<close>
declare sorted.Nil [code_pred_intro]
sorted_single [code_pred_intro]
@@ -69,15 +69,15 @@
case (Cons y zs)
show ?thesis
proof (rule 3)
- from Cons `xa = x # xs` show "xa = x # y # zs" by simp
- with `sorted xa` show "x \<le> y" and "sorted (y # zs)" by simp_all
+ from Cons \<open>xa = x # xs\<close> show "xa = x # y # zs" by simp
+ with \<open>sorted xa\<close> show "x \<le> y" and "sorted (y # zs)" by simp_all
qed
qed
qed
qed
thm sorted.equation
-section {* Specialisation in POPLmark theory *}
+section \<open>Specialisation in POPLmark theory\<close>
notation
Some ("\<lfloor>_\<rfloor>")
@@ -250,11 +250,11 @@
values 6 "{(E, t, T). typing E t T}"
-subsection {* Higher-order predicate *}
+subsection \<open>Higher-order predicate\<close>
code_pred [inductify] mapB .
-subsection {* Multiple instances *}
+subsection \<open>Multiple instances\<close>
inductive subtype_refl' where
"\<Gamma> \<turnstile> t : T ==> \<not> (\<Gamma> \<turnstile> T <: T) ==> subtype_refl' t T"
--- a/src/HOL/Probability/Borel_Space.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Probability/Borel_Space.thy Thu May 26 17:51:22 2016 +0200
@@ -156,7 +156,7 @@
shows "countable {a\<in>A. \<not> continuous (at a within A) f}"
proof -
have mono: "\<And>x y. x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> x \<le> y \<Longrightarrow> f x \<le> f y"
- using `mono_on f A` by (simp add: mono_on_def)
+ using \<open>mono_on f A\<close> by (simp add: mono_on_def)
have "\<forall>a \<in> {a\<in>A. \<not> continuous (at a within A) f}. \<exists>q :: nat \<times> rat.
(fst q = 0 \<and> of_rat (snd q) < f a \<and> (\<forall>x \<in> A. x < a \<longrightarrow> f x < of_rat (snd q))) \<or>
(fst q = 1 \<and> of_rat (snd q) > f a \<and> (\<forall>x \<in> A. x > a \<longrightarrow> f x > of_rat (snd q)))"
@@ -224,7 +224,7 @@
shows "countable {a\<in>A. \<not>isCont f a}"
proof -
have "{a\<in>A. \<not>isCont f a} = {a\<in>A. \<not>(continuous (at a within A) f)}"
- by (auto simp add: continuous_within_open [OF _ `open A`])
+ by (auto simp add: continuous_within_open [OF _ \<open>open A\<close>])
thus ?thesis
apply (elim ssubst)
by (rule mono_on_ctble_discont, rule assms)
--- a/src/HOL/Probability/Characteristic_Functions.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Probability/Characteristic_Functions.thy Thu May 26 17:51:22 2016 +0200
@@ -270,7 +270,7 @@
qed
lemma iexp_approx2: "cmod (iexp x - (\<Sum>k \<le> n. (ii * x)^k / fact k)) \<le> 2 * \<bar>x\<bar>^n / fact n"
-proof (induction n) -- \<open>really cases\<close>
+proof (induction n) \<comment> \<open>really cases\<close>
case (Suc n)
have *: "\<And>a b. interval_lebesgue_integrable lborel a b f \<Longrightarrow> interval_lebesgue_integrable lborel a b g \<Longrightarrow>
\<bar>LBINT s=a..b. f s\<bar> \<le> \<bar>LBINT s=a..b. g s\<bar>"
--- a/src/HOL/Probability/Distribution_Functions.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Probability/Distribution_Functions.thy Thu May 26 17:51:22 2016 +0200
@@ -26,7 +26,7 @@
(metis le_less_trans minus_minus neg_less_iff_less not_le real_arch_simple
of_nat_0_le_iff reals_Archimedean2)
-subsection {* Properties of cdf's *}
+subsection \<open>Properties of cdf's\<close>
definition
cdf :: "real measure \<Rightarrow> real \<Rightarrow> real"
@@ -105,7 +105,7 @@
proof (rule tendsto_at_right_sequentially[where b="a + 1"])
fix f :: "nat \<Rightarrow> real" and x assume f: "decseq f" "f \<longlonglongrightarrow> a"
then have "(\<lambda>n. cdf M (f n)) \<longlonglongrightarrow> measure M (\<Inter>i. {.. f i})"
- using `decseq f` unfolding cdf_def
+ using \<open>decseq f\<close> unfolding cdf_def
by (intro finite_Lim_measure_decseq) (auto simp: decseq_def)
also have "(\<Inter>i. {.. f i}) = {.. a}"
using decseq_le[OF f] by (auto intro: order_trans LIMSEQ_le_const[OF f(2)])
@@ -117,7 +117,7 @@
proof (rule tendsto_at_left_sequentially[of "a - 1"])
fix f :: "nat \<Rightarrow> real" and x assume f: "incseq f" "f \<longlonglongrightarrow> a" "\<And>x. f x < a" "\<And>x. a - 1 < f x"
then have "(\<lambda>n. cdf M (f n)) \<longlonglongrightarrow> measure M (\<Union>i. {.. f i})"
- using `incseq f` unfolding cdf_def
+ using \<open>incseq f\<close> unfolding cdf_def
by (intro finite_Lim_measure_incseq) (auto simp: incseq_def)
also have "(\<Union>i. {.. f i}) = {..<a}"
by (auto dest!: order_tendstoD(1)[OF f(2)] eventually_happens'[OF sequentially_bot]
@@ -168,14 +168,14 @@
"random_variable borel X \<Longrightarrow> real_distribution (distr M borel X)"
unfolding real_distribution_def real_distribution_axioms_def by (auto intro!: prob_space_distr)
-subsection {* uniqueness *}
+subsection \<open>uniqueness\<close>
lemma (in real_distribution) emeasure_Ioc:
assumes "a \<le> b" shows "emeasure M {a <.. b} = cdf M b - cdf M a"
proof -
have "{a <.. b} = {..b} - {..a}"
by auto
- with `a \<le> b` show ?thesis
+ with \<open>a \<le> b\<close> show ?thesis
by (simp add: emeasure_eq_measure finite_measure_Diff cdf_def)
qed
--- a/src/HOL/Probability/Levy.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Probability/Levy.thy Thu May 26 17:51:22 2016 +0200
@@ -28,7 +28,7 @@
(iexp (-(t * a)) - (1 + ii * -(t * a))) / (ii * t) -
(iexp (-(t * b)) - (1 + ii * -(t * b))) / (ii * t))"
(is "_ = cmod (?one / (ii * t) - ?two / (ii * t))")
- using `t \<noteq> 0` by (intro arg_cong[where f=norm]) (simp add: field_simps)
+ using \<open>t \<noteq> 0\<close> by (intro arg_cong[where f=norm]) (simp add: field_simps)
also have "\<dots> \<le> cmod (?one / (ii * t)) + cmod (?two / (ii * t))"
by (rule norm_triangle_ineq4)
also have "cmod (?one / (ii * t)) = cmod ?one / abs t"
@@ -45,8 +45,8 @@
using iexp_approx1 [of "-(t * b)" 1] apply (simp add: field_simps eval_nat_numeral)
by force
also have "\<dots> = a^2 / 2 * abs t + b^2 / 2 * abs t"
- using `t \<noteq> 0` apply (case_tac "t \<ge> 0", simp add: field_simps power2_eq_square)
- using `t \<noteq> 0` by (subst (1 2) abs_of_neg, auto simp add: field_simps power2_eq_square)
+ using \<open>t \<noteq> 0\<close> apply (case_tac "t \<ge> 0", simp add: field_simps power2_eq_square)
+ using \<open>t \<noteq> 0\<close> by (subst (1 2) abs_of_neg, auto simp add: field_simps power2_eq_square)
finally show "cmod (?F t - (b - a)) \<le> a^2 / 2 * abs t + b^2 / 2 * abs t" .
qed
show ?thesis
@@ -63,9 +63,9 @@
shows "cmod ((iexp (t * b) - iexp (t * a)) / (ii * t)) \<le> b - a" (is "?F \<le> _")
proof -
have "?F = cmod (iexp (t * a) * (iexp (t * (b - a)) - 1) / (ii * t))"
- using `t \<noteq> 0` by (intro arg_cong[where f=norm]) (simp add: field_simps exp_diff exp_minus)
+ using \<open>t \<noteq> 0\<close> by (intro arg_cong[where f=norm]) (simp add: field_simps exp_diff exp_minus)
also have "\<dots> = cmod (iexp (t * (b - a)) - 1) / abs t"
- unfolding norm_divide norm_mult norm_exp_ii_times using `t \<noteq> 0`
+ unfolding norm_divide norm_mult norm_exp_ii_times using \<open>t \<noteq> 0\<close>
by (simp add: complex_eq_iff norm_mult)
also have "\<dots> \<le> abs (t * (b - a)) / abs t"
using iexp_approx1 [of "t * (b - a)" 0]
@@ -100,7 +100,7 @@
apply (auto intro!: AE_I [of _ _ "{0}"] simp: assms)
done
have "(CLBINT t. ?f' (t, x)) = (CLBINT t=-T..T. ?f t x)"
- using `T \<ge> 0` by (simp add: interval_lebesgue_integral_def)
+ using \<open>T \<ge> 0\<close> by (simp add: interval_lebesgue_integral_def)
also have "\<dots> = (CLBINT t=-T..(0 :: real). ?f t x) + (CLBINT t=(0 :: real)..T. ?f t x)"
(is "_ = _ + ?t")
using 1 by (intro interval_integral_sum[symmetric]) (simp add: min_absorb1 max_absorb2 \<open>T \<ge> 0\<close>)
@@ -111,10 +111,10 @@
by (intro interval_lebesgue_integral_add(2) [symmetric] interval_integrable_mirror[THEN iffD2]) simp_all
also have "\<dots> = (CLBINT t=(0::real)..T. ((iexp(t * (x - a)) - iexp (-(t * (x - a)))) -
(iexp(t * (x - b)) - iexp (-(t * (x - b))))) / (ii * t))"
- using `T \<ge> 0` by (intro interval_integral_cong) (auto simp add: divide_simps)
+ using \<open>T \<ge> 0\<close> by (intro interval_integral_cong) (auto simp add: divide_simps)
also have "\<dots> = (CLBINT t=(0::real)..T. complex_of_real(
2 * (sin (t * (x - a)) / t) - 2 * (sin (t * (x - b)) / t)))"
- using `T \<ge> 0`
+ using \<open>T \<ge> 0\<close>
apply (intro interval_integral_cong)
apply (simp add: field_simps cis.ctr Im_divide Re_divide Im_exp Re_exp complex_eq_iff)
unfolding minus_diff_eq[symmetric, of "y * x" "y * a" for y a] sin_minus cos_minus
@@ -128,14 +128,14 @@
apply (subst interval_lebesgue_integral_diff)
apply (rule interval_lebesgue_integrable_mult_right, rule integrable_sinc')+
apply (subst interval_lebesgue_integral_mult_right)+
- apply (simp add: zero_ereal_def[symmetric] LBINT_I0c_sin_scale_divide[OF `T \<ge> 0`])
+ apply (simp add: zero_ereal_def[symmetric] LBINT_I0c_sin_scale_divide[OF \<open>T \<ge> 0\<close>])
done
finally have "(CLBINT t. ?f' (t, x)) =
2 * (sgn (x - a) * Si (T * abs (x - a)) - sgn (x - b) * Si (T * abs (x - b)))" .
} note main_eq = this
have "(CLBINT t=-T..T. ?F t * \<phi> t) =
(CLBINT t. (CLINT x | M. ?F t * iexp (t * x) * indicator {-T<..<T} t))"
- using `T \<ge> 0` unfolding \<phi>_def char_def interval_lebesgue_integral_def
+ using \<open>T \<ge> 0\<close> unfolding \<phi>_def char_def interval_lebesgue_integral_def
by (auto split: split_indicator intro!: integral_cong)
also have "\<dots> = (CLBINT t. (CLINT x | M. ?f' (t, x)))"
by (auto intro!: integral_cong simp: field_simps exp_diff exp_minus split: split_indicator)
@@ -146,7 +146,7 @@
by (subst emeasure_pair_measure_Times)
(auto simp: ennreal_mult_less_top not_less top_unique)
show "AE x\<in>{- T<..<T} \<times> space M in lborel \<Otimes>\<^sub>M M. cmod (case x of (t, x) \<Rightarrow> ?f' (t, x)) \<le> b - a"
- using Levy_Inversion_aux2[of "x - b" "x - a" for x] `a \<le> b`
+ using Levy_Inversion_aux2[of "x - b" "x - a" for x] \<open>a \<le> b\<close>
by (intro AE_I [of _ _ "{0} \<times> UNIV"]) (force simp: emeasure_pair_measure_Times)+
qed (auto split: split_indicator split_indicator_asm)
also have "\<dots> = (CLINT x | M. (complex_of_real (2 * (sgn (x - a) *
@@ -178,7 +178,7 @@
then show "\<And>n. AE x in M. norm (2 * ?S n x) \<le> 4 * B"
by auto
have "AE x in M. x \<noteq> a" "AE x in M. x \<noteq> b"
- using prob_eq_0[of "{a}"] prob_eq_0[of "{b}"] `\<mu> {a} = 0` `\<mu> {b} = 0` by (auto simp: \<mu>_def)
+ using prob_eq_0[of "{a}"] prob_eq_0[of "{b}"] \<open>\<mu> {a} = 0\<close> \<open>\<mu> {b} = 0\<close> by (auto simp: \<mu>_def)
then show "AE x in M. (\<lambda>n. 2 * ?S n x) \<longlonglongrightarrow> 2 * pi * indicator {a<..b} x"
proof eventually_elim
fix x assume x: "x \<noteq> a" "x \<noteq> b"
@@ -190,7 +190,7 @@
also have "(\<lambda>n. 2 * (sgn (x - a) * Si (\<bar>x - a\<bar> * n) - sgn (x - b) * Si (\<bar>x - b\<bar> * n))) = (\<lambda>n. 2 * ?S n x)"
by (auto simp: ac_simps)
also have "2 * (sgn (x - a) * (pi / 2) - sgn (x - b) * (pi / 2)) = 2 * pi * indicator {a<..b} x"
- using x `a \<le> b` by (auto split: split_indicator)
+ using x \<open>a \<le> b\<close> by (auto split: split_indicator)
finally show "(\<lambda>n. 2 * ?S n x) \<longlonglongrightarrow> 2 * pi * indicator {a<..b} x" .
qed
qed simp_all
@@ -226,7 +226,7 @@
{ fix \<epsilon> :: real
assume "\<epsilon> > 0"
- from `\<epsilon> > 0` `(cdf M1 \<longlongrightarrow> 0) at_bot` `(cdf M2 \<longlongrightarrow> 0) at_bot`
+ from \<open>\<epsilon> > 0\<close> \<open>(cdf M1 \<longlongrightarrow> 0) at_bot\<close> \<open>(cdf M2 \<longlongrightarrow> 0) at_bot\<close>
have "eventually (\<lambda>y. \<bar>cdf M1 y\<bar> < \<epsilon> / 4 \<and> \<bar>cdf M2 y\<bar> < \<epsilon> / 4 \<and> y \<le> x) at_bot"
by (simp only: tendsto_iff dist_real_def diff_0_right eventually_conj eventually_le_at_bot)
then obtain M where "\<And>y. y \<le> M \<Longrightarrow> \<bar>cdf M1 y\<bar> < \<epsilon> / 4" "\<And>y. y \<le> M \<Longrightarrow> \<bar>cdf M2 y\<bar> < \<epsilon> / 4" "M \<le> x"
@@ -235,7 +235,7 @@
"measure M1 {a} = 0" "measure M2 {a} = 0" "a < M" "a \<le> x" "\<bar>cdf M1 a\<bar> < \<epsilon> / 4" "\<bar>cdf M2 a\<bar> < \<epsilon> / 4"
by auto
- from `\<epsilon> > 0` `(cdf M1 \<longlongrightarrow> cdf M1 x) (at_right x)` `(cdf M2 \<longlongrightarrow> cdf M2 x) (at_right x)`
+ from \<open>\<epsilon> > 0\<close> \<open>(cdf M1 \<longlongrightarrow> cdf M1 x) (at_right x)\<close> \<open>(cdf M2 \<longlongrightarrow> cdf M2 x) (at_right x)\<close>
have "eventually (\<lambda>y. \<bar>cdf M1 y - cdf M1 x\<bar> < \<epsilon> / 4 \<and> \<bar>cdf M2 y - cdf M2 x\<bar> < \<epsilon> / 4 \<and> x < y) (at_right x)"
by (simp only: tendsto_iff dist_real_def eventually_conj eventually_at_right_less)
then obtain N where "N > x" "\<And>y. x < y \<Longrightarrow> y < N \<Longrightarrow> \<bar>cdf M1 y - cdf M1 x\<bar> < \<epsilon> / 4"
@@ -244,16 +244,16 @@
with open_minus_countable[OF count, of "{x <..< N}"] obtain b where "x < b" "b < N"
"measure M1 {b} = 0" "measure M2 {b} = 0" "\<bar>cdf M2 x - cdf M2 b\<bar> < \<epsilon> / 4" "\<bar>cdf M1 x - cdf M1 b\<bar> < \<epsilon> / 4"
by (auto simp: abs_minus_commute)
- from `a \<le> x` `x < b` have "a < b" "a \<le> b" by auto
+ from \<open>a \<le> x\<close> \<open>x < b\<close> have "a < b" "a \<le> b" by auto
- from `char M1 = char M2`
- M1.Levy_Inversion [OF `a \<le> b` `measure M1 {a} = 0` `measure M1 {b} = 0`]
- M2.Levy_Inversion [OF `a \<le> b` `measure M2 {a} = 0` `measure M2 {b} = 0`]
+ from \<open>char M1 = char M2\<close>
+ M1.Levy_Inversion [OF \<open>a \<le> b\<close> \<open>measure M1 {a} = 0\<close> \<open>measure M1 {b} = 0\<close>]
+ M2.Levy_Inversion [OF \<open>a \<le> b\<close> \<open>measure M2 {a} = 0\<close> \<open>measure M2 {b} = 0\<close>]
have "complex_of_real (measure M1 {a<..b}) = complex_of_real (measure M2 {a<..b})"
by (intro LIMSEQ_unique) auto
then have "measure M1 {a<..b} = measure M2 {a<..b}" by auto
then have *: "cdf M1 b - cdf M1 a = cdf M2 b - cdf M2 a"
- unfolding M1.cdf_diff_eq [OF `a < b`] M2.cdf_diff_eq [OF `a < b`] .
+ unfolding M1.cdf_diff_eq [OF \<open>a < b\<close>] M2.cdf_diff_eq [OF \<open>a < b\<close>] .
have "abs (cdf M1 x - (cdf M1 b - cdf M1 a)) = abs (cdf M1 x - cdf M1 b + cdf M1 a)"
by simp
@@ -281,7 +281,7 @@
by (metis abs_le_zero_iff dense_ge eq_iff_diff_eq_0)
qed
thus ?thesis
- by (rule cdf_unique [OF `real_distribution M1` `real_distribution M2`])
+ by (rule cdf_unique [OF \<open>real_distribution M1\<close> \<open>real_distribution M2\<close>])
qed
@@ -311,7 +311,7 @@
hence "(CLBINT t:{-u..u}. 1 - iexp (t * x)) = (CLBINT t=-u..u. 1 - iexp (t * x))"
by (subst interval_integral_Icc, auto)
also have "\<dots> = (CLBINT t=-u..0. 1 - iexp (t * x)) + (CLBINT t=0..u. 1 - iexp (t * x))"
- using `u > 0`
+ using \<open>u > 0\<close>
apply (subst interval_integral_sum)
apply (simp add: min_absorb1 min_absorb2 max_absorb1 max_absorb2)
apply (rule interval_integrable_isCont)
@@ -329,7 +329,7 @@
also have "\<dots> = 2 * u - 2 * sin (u * x) / x"
by (subst interval_lebesgue_integral_diff)
(auto intro!: interval_integrable_isCont
- simp: interval_lebesgue_integral_of_real integral_cos [OF `x \<noteq> 0`] mult.commute[of _ x])
+ simp: interval_lebesgue_integral_of_real integral_cos [OF \<open>x \<noteq> 0\<close>] mult.commute[of _ x])
finally show "(CLBINT t:{-u..u}. 1 - iexp (t * x)) = 2 * (u - sin (u * x) / x)"
by (simp add: field_simps)
qed
@@ -345,7 +345,7 @@
have Mn2 [simp]: "\<And>x. complex_integrable (M n) (\<lambda>t. exp (\<i> * complex_of_real (x * t)))"
by (rule Mn.integrable_const_bound [where B = 1], auto)
have Mn3: "set_integrable (M n \<Otimes>\<^sub>M lborel) (UNIV \<times> {- u..u}) (\<lambda>a. 1 - exp (\<i> * complex_of_real (snd a * fst a)))"
- using `0 < u`
+ using \<open>0 < u\<close>
by (intro integrableI_bounded_set_indicator [where B="2"])
(auto simp: lborel.emeasure_pair_measure_Times ennreal_mult_less_top not_less top_unique
split: split_indicator
@@ -358,7 +358,7 @@
also have "\<dots> = (CLINT x | M n. (CLBINT t:{-u..u}. 1 - iexp (t * x)))"
using Mn3 by (subst P.Fubini_integral) (auto simp: indicator_times split_beta')
also have "\<dots> = (CLINT x | M n. (if x = 0 then 0 else 2 * (u - sin (u * x) / x)))"
- using `u > 0` by (intro integral_cong, auto simp add: * simp del: of_real_mult)
+ using \<open>u > 0\<close> by (intro integral_cong, auto simp add: * simp del: of_real_mult)
also have "\<dots> = (LINT x | M n. (if x = 0 then 0 else 2 * (u - sin (u * x) / x)))"
by (rule integral_complex_of_real)
finally have "Re (CLBINT t:{-u..u}. 1 - char (M n) t) =
@@ -368,7 +368,7 @@
have "complex_integrable (M n) (\<lambda>x. CLBINT t:{-u..u}. 1 - iexp (snd (x, t) * fst (x, t)))"
using Mn3 by (intro P.integrable_fst) (simp add: indicator_times split_beta')
hence "complex_integrable (M n) (\<lambda>x. if x = 0 then 0 else 2 * (u - sin (u * x) / x))"
- using `u > 0` by (subst integrable_cong) (auto simp add: * simp del: of_real_mult)
+ using \<open>u > 0\<close> by (subst integrable_cong) (auto simp add: * simp del: of_real_mult)
hence **: "integrable (M n) (\<lambda>x. if x = 0 then 0 else 2 * (u - sin (u * x) / x))"
unfolding complex_of_real_integrable_eq .
have "2 * sin x \<le> x" if "2 \<le> x" for x :: real
@@ -378,7 +378,7 @@
moreover have "x < 0 \<Longrightarrow> x \<le> sin x" for x :: real
using sin_x_le_x[of "-x"] by simp
ultimately show ?thesis
- using `u > 0`
+ using \<open>u > 0\<close>
by (intro integral_mono [OF _ **])
(auto simp: divide_simps sin_x_le_x mult.commute[of u] mult_neg_pos top_unique less_top[symmetric]
split: split_indicator)
@@ -397,7 +397,7 @@
hence "\<exists>d>0. \<forall>t. abs t < d \<longrightarrow> cmod (char M' t - 1) < \<epsilon> / 4"
apply (subst (asm) continuous_at_eps_delta)
apply (drule_tac x = "\<epsilon> / 4" in spec)
- using `\<epsilon> > 0` by (auto simp add: dist_real_def dist_complex_def M'.char_zero)
+ using \<open>\<epsilon> > 0\<close> by (auto simp add: dist_real_def dist_complex_def M'.char_zero)
then obtain d where "d > 0 \<and> (\<forall>t. (abs t < d \<longrightarrow> cmod (char M' t - 1) < \<epsilon> / 4))" ..
hence d0: "d > 0" and d1: "\<And>t. abs t < d \<Longrightarrow> cmod (char M' t - 1) < \<epsilon> / 4" by auto
have 1: "\<And>x. cmod (1 - char M' x) \<le> 2"
@@ -433,7 +433,7 @@
done
hence "eventually (\<lambda>n. cmod ((CLBINT t:{-d/2..d/2}. 1 - char (M n) t) -
(CLBINT t:{-d/2..d/2}. 1 - char M' t)) < d * \<epsilon> / 4) sequentially"
- using d0 `\<epsilon> > 0` apply (subst (asm) tendsto_iff)
+ using d0 \<open>\<epsilon> > 0\<close> apply (subst (asm) tendsto_iff)
by (subst (asm) dist_complex_def, drule spec, erule mp, auto)
hence "\<exists>N. \<forall>n \<ge> N. cmod ((CLBINT t:{-d/2..d/2}. 1 - char (M n) t) -
(CLBINT t:{-d/2..d/2}. 1 - char M' t)) < d * \<epsilon> / 4" by (simp add: eventually_sequentially)
@@ -449,7 +449,7 @@
(CLBINT t:{-d/2..d/2}. 1 - char M' t)) + cmod(CLBINT t:{-d/2..d/2}. 1 - char M' t)"
by (rule norm_triangle_ineq)
also have "\<dots> < d * \<epsilon> / 4 + d * \<epsilon> / 4"
- by (rule add_less_le_mono [OF N [OF `n \<ge> N`] bound])
+ by (rule add_less_le_mono [OF N [OF \<open>n \<ge> N\<close>] bound])
also have "\<dots> = d * \<epsilon> / 2" by auto
finally have "cmod (CLBINT t:{-d/2..d/2}. 1 - char (M n) t) < d * \<epsilon> / 2" .
hence "d * \<epsilon> / 2 > Re (CLBINT t:{-d/2..d/2}. 1 - char (M n) t)"
@@ -458,7 +458,7 @@
also have "?lhs \<ge> (d / 2) * measure (M n) {x. abs x \<ge> 2 / (d / 2)}"
using d0 by (intro main_bound, simp)
finally (xtrans) have "d * \<epsilon> / 2 > (d / 2) * measure (M n) {x. abs x \<ge> 2 / (d / 2)}" .
- with d0 `\<epsilon> > 0` have "\<epsilon> > measure (M n) {x. abs x \<ge> 2 / (d / 2)}" by (simp add: field_simps)
+ with d0 \<open>\<epsilon> > 0\<close> have "\<epsilon> > measure (M n) {x. abs x \<ge> 2 / (d / 2)}" by (simp add: field_simps)
hence "\<epsilon> > 1 - measure (M n) (UNIV - {x. abs x \<ge> 2 / (d / 2)})"
apply (subst Mn.borel_UNIV [symmetric])
by (subst Mn.prob_compl, auto)
@@ -484,7 +484,7 @@
using Mn.prob_space unfolding * Mn.borel_UNIV by simp
hence "eventually (\<lambda>k. measure (M n) {- real k<..real k} > 1 - \<epsilon>) sequentially"
apply (elim order_tendstoD (1))
- using `\<epsilon> > 0` by auto
+ using \<open>\<epsilon> > 0\<close> by auto
} note 7 = this
{ fix n :: nat
have "eventually (\<lambda>k. \<forall>m < n. measure (M m) {- real k<..real k} > 1 - \<epsilon>) sequentially"
@@ -533,7 +533,7 @@
have 5: "\<And>t. (\<lambda>n. char ((M \<circ> s) n) t) \<longlonglongrightarrow> char M' t"
by (subst 4, rule LIMSEQ_subseq_LIMSEQ [OF _ s], rule assms)
hence "char \<nu> = char M'" by (intro ext, intro LIMSEQ_unique [OF 3 5])
- hence "\<nu> = M'" by (rule Levy_uniqueness [OF * `real_distribution M'`])
+ hence "\<nu> = M'" by (rule Levy_uniqueness [OF * \<open>real_distribution M'\<close>])
thus "weak_conv_m (M \<circ> s) M'"
by (elim subst) (rule nu)
qed
--- a/src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy Thu May 26 17:51:22 2016 +0200
@@ -1648,7 +1648,7 @@
using sums_If_finite[of "\<lambda>r. r < f x" "\<lambda>_. 1 :: ennreal"]
by (cases "f x") (simp_all add: sums_def of_nat_tendsto_top_ennreal)
also have "(\<lambda>i. (if i < f x then 1 else 0)) = (\<lambda>i. indicator (F i) x)"
- using `x \<in> space M` by (simp add: one_ennreal_def F_def fun_eq_iff)
+ using \<open>x \<in> space M\<close> by (simp add: one_ennreal_def F_def fun_eq_iff)
finally have "ennreal_of_enat (f x) = (\<Sum>i. indicator (F i) x)"
by (simp add: sums_iff) }
then have "(\<integral>\<^sup>+x. ennreal_of_enat (f x) \<partial>M) = (\<integral>\<^sup>+x. (\<Sum>i. indicator (F i) x) \<partial>M)"
--- a/src/HOL/Probability/Sigma_Algebra.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Probability/Sigma_Algebra.thy Thu May 26 17:51:22 2016 +0200
@@ -542,7 +542,7 @@
using from_nat_into [of B] range_from_nat_into [of B] sigma_sets.Union [of "from_nat_into B" X A]
apply simp
apply auto
- apply (metis Sup_bot_conv(1) Union_empty `\<lbrakk>B \<noteq> {}; countable B\<rbrakk> \<Longrightarrow> range (from_nat_into B) = B`)
+ apply (metis Sup_bot_conv(1) Union_empty \<open>\<lbrakk>B \<noteq> {}; countable B\<rbrakk> \<Longrightarrow> range (from_nat_into B) = B\<close>)
done
lemma (in sigma_algebra) sigma_sets_eq:
--- a/src/HOL/Probability/Sinc_Integral.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Probability/Sinc_Integral.thy Thu May 26 17:51:22 2016 +0200
@@ -13,7 +13,7 @@
text \<open> Naming convention
The theorem name consists of the following parts:
- \<^item> Kind of integral: @{text has_bochner_integral} / @{text integrable} / @{text LBINT}
+ \<^item> Kind of integral: \<open>has_bochner_integral\<close> / \<open>integrable\<close> / \<open>LBINT\<close>
\<^item> Interval: Interval (0 / infinity / open / closed) (infinity / open / closed)
\<^item> Name of the occurring constants: power, exp, m (for minus), scale, sin, $\ldots$
\<close>
@@ -273,7 +273,7 @@
proof eventually_elim
fix t :: real assume "t \<ge> 0"
have "Si t = LBINT x=0..t. sin x * (LBINT u=0..\<infinity>. exp (-(u * x)))"
- unfolding Si_def using `0 \<le> t`
+ unfolding Si_def using \<open>0 \<le> t\<close>
by (intro interval_integral_discrete_difference[where X="{0}"]) (auto simp: LBINT_I0i_exp_mscale)
also have "\<dots> = LBINT x. (LBINT u=ereal 0..\<infinity>. indicator {0 <..< t} x *\<^sub>R sin x * exp (-(u * x)))"
using \<open>0\<le>t\<close> by (simp add: zero_ereal_def interval_lebesgue_integral_le_eq mult_ac)
@@ -369,7 +369,7 @@
by simp
hence "LBINT x. indicator {0<..<T} x * sin (x * \<theta>) / x =
LBINT x. indicator {0<..<T * \<theta>} x * sin x / x"
- using assms `0 < \<theta>` unfolding interval_lebesgue_integral_def einterval_eq zero_ereal_def
+ using assms \<open>0 < \<theta>\<close> unfolding interval_lebesgue_integral_def einterval_eq zero_ereal_def
by (auto simp: ac_simps)
} note aux1 = this
{ assume "0 > \<theta>"
@@ -388,7 +388,7 @@
by simp
hence "LBINT x. indicator {0<..<T} x * sin (x * \<theta>) / x =
- (LBINT x. indicator {0<..<- (T * \<theta>)} x * sin x / x)"
- using assms `0 > \<theta>` unfolding interval_lebesgue_integral_def einterval_eq zero_ereal_def
+ using assms \<open>0 > \<theta>\<close> unfolding interval_lebesgue_integral_def einterval_eq zero_ereal_def
by (auto simp add: field_simps mult_le_0_iff split: if_split_asm)
} note aux2 = this
show ?thesis
--- a/src/HOL/Prolog/Func.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Prolog/Func.thy Thu May 26 17:51:22 2016 +0200
@@ -2,7 +2,7 @@
Author: David von Oheimb (based on a lecture on Lambda Prolog by Nadathur)
*)
-section {* Untyped functional language, with call by value semantics *}
+section \<open>Untyped functional language, with call by value semantics\<close>
theory Func
imports HOHH
--- a/src/HOL/Prolog/HOHH.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Prolog/HOHH.thy Thu May 26 17:51:22 2016 +0200
@@ -2,7 +2,7 @@
Author: David von Oheimb (based on a lecture on Lambda Prolog by Nadathur)
*)
-section {* Higher-order hereditary Harrop formulas *}
+section \<open>Higher-order hereditary Harrop formulas\<close>
theory HOHH
imports HOL
@@ -11,11 +11,11 @@
ML_file "prolog.ML"
method_setup ptac =
- {* Attrib.thms >> (fn thms => fn ctxt => SIMPLE_METHOD' (Prolog.ptac ctxt thms)) *}
+ \<open>Attrib.thms >> (fn thms => fn ctxt => SIMPLE_METHOD' (Prolog.ptac ctxt thms))\<close>
"basic Lambda Prolog interpreter"
method_setup prolog =
- {* Attrib.thms >> (fn thms => fn ctxt => SIMPLE_METHOD (Prolog.prolog_tac ctxt thms)) *}
+ \<open>Attrib.thms >> (fn thms => fn ctxt => SIMPLE_METHOD (Prolog.prolog_tac ctxt thms))\<close>
"Lambda Prolog interpreter"
consts
--- a/src/HOL/Prolog/Test.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Prolog/Test.thy Thu May 26 17:51:22 2016 +0200
@@ -2,7 +2,7 @@
Author: David von Oheimb (based on a lecture on Lambda Prolog by Nadathur)
*)
-section {* Basic examples *}
+section \<open>Basic examples\<close>
theory Test
imports HOHH
--- a/src/HOL/Prolog/Type.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Prolog/Type.thy Thu May 26 17:51:22 2016 +0200
@@ -2,7 +2,7 @@
Author: David von Oheimb (based on a lecture on Lambda Prolog by Nadathur)
*)
-section {* Type inference *}
+section \<open>Type inference\<close>
theory Type
imports Func
--- a/src/HOL/Quickcheck_Examples/Completeness.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Quickcheck_Examples/Completeness.thy Thu May 26 17:51:22 2016 +0200
@@ -3,13 +3,13 @@
Copyright 2012 TU Muenchen
*)
-section {* Proving completeness of exhaustive generators *}
+section \<open>Proving completeness of exhaustive generators\<close>
theory Completeness
imports Main
begin
-subsection {* Preliminaries *}
+subsection \<open>Preliminaries\<close>
primrec is_some
where
@@ -22,7 +22,7 @@
setup Exhaustive_Generators.setup_exhaustive_datatype_interpretation
-subsection {* Defining the size of values *}
+subsection \<open>Defining the size of values\<close>
hide_const size
@@ -74,7 +74,7 @@
end
-subsection {* Completeness *}
+subsection \<open>Completeness\<close>
class complete = exhaustive + size +
assumes
@@ -89,7 +89,7 @@
obtains v where "size (v :: 'a :: complete) \<le> n" "is_some (f v)"
using assms by (metis complete)
-subsubsection {* Instance Proofs *}
+subsubsection \<open>Instance Proofs\<close>
declare exhaustive_int'.simps [simp del]
--- a/src/HOL/Quickcheck_Examples/Hotel_Example.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Quickcheck_Examples/Hotel_Example.thy Thu May 26 17:51:22 2016 +0200
@@ -83,7 +83,7 @@
no_Check_in (s\<^sub>3 @ s\<^sub>2) r \<and> isin (s\<^sub>2 @ [Check_in g r c] @ s\<^sub>1) r = {})"
-section {* Some setup *}
+section \<open>Some setup\<close>
lemma issued_nil: "issued [] = {Key0}"
by (auto simp add: initk_def)
@@ -91,10 +91,10 @@
lemmas issued_simps[code] = issued_nil issued.simps(2)
-setup {* Predicate_Compile_Data.ignore_consts [@{const_name Set.member},
+setup \<open>Predicate_Compile_Data.ignore_consts [@{const_name Set.member},
@{const_name "issued"}, @{const_name "cards"}, @{const_name "isin"},
- @{const_name Collect}, @{const_name insert}] *}
-ML_val {* Core_Data.force_modes_and_compilations *}
+ @{const_name Collect}, @{const_name insert}]\<close>
+ML_val \<open>Core_Data.force_modes_and_compilations\<close>
fun find_first :: "('a => 'b option) => 'a list => 'b option"
where
@@ -116,7 +116,7 @@
axiomatization neg_cps_of_set :: "'a set => ('a Quickcheck_Exhaustive.unknown => term list Quickcheck_Exhaustive.three_valued) => natural => term list Quickcheck_Exhaustive.three_valued"
where neg_cps_of_set_code [code]: "neg_cps_of_set (set xs) f i = find_first' f xs"
-setup {*
+setup \<open>
let
val Fun = Predicate_Compile_Aux.Fun
val Input = Predicate_Compile_Aux.Input
@@ -137,8 +137,8 @@
Core_Data.force_modes_and_compilations @{const_name Set.member}
[(oi, (of_set, false)), (ii, (member, false))]
end
-*}
-section {* Property *}
+\<close>
+section \<open>Property\<close>
lemma "\<lbrakk> hotel s; g \<in> isin s r \<rbrakk> \<Longrightarrow> owns s r = Some g"
quickcheck[tester = exhaustive, size = 6, expect = counterexample]
@@ -150,7 +150,7 @@
quickcheck[smart_exhaustive, depth = 10, allow_function_inversion, expect = counterexample]
oops
-section {* Refinement *}
+section \<open>Refinement\<close>
fun split_list
where
--- a/src/HOL/Quickcheck_Examples/Quickcheck_Examples.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Quickcheck_Examples/Quickcheck_Examples.thy Thu May 26 17:51:22 2016 +0200
@@ -3,13 +3,13 @@
Copyright 2004 - 2010 TU Muenchen
*)
-section {* Examples for the 'quickcheck' command *}
+section \<open>Examples for the 'quickcheck' command\<close>
theory Quickcheck_Examples
imports Complex_Main "~~/src/HOL/Library/Dlist" "~~/src/HOL/Library/DAList_Multiset"
begin
-text {*
+text \<open>
The 'quickcheck' command allows to find counterexamples by evaluating
formulae.
Currently, there are two different exploration schemes:
@@ -19,11 +19,11 @@
quickcheck can handle quantifiers on finite universes.
-*}
+\<close>
declare [[quickcheck_timeout = 3600]]
-subsection {* Lists *}
+subsection \<open>Lists\<close>
theorem "map g (map f xs) = map (g o f) xs"
quickcheck[random, expect = no_counterexample]
@@ -61,7 +61,7 @@
oops
-text {* An example involving functions inside other data structures *}
+text \<open>An example involving functions inside other data structures\<close>
primrec app :: "('a \<Rightarrow> 'a) list \<Rightarrow> 'a \<Rightarrow> 'a" where
"app [] x = x"
@@ -85,9 +85,9 @@
"del1 a [] = []"
| "del1 a (x#xs) = (if (x=a) then xs else (x#del1 a xs))"
-text {* A lemma, you'd think to be true from our experience with delAll *}
+text \<open>A lemma, you'd think to be true from our experience with delAll\<close>
lemma "Suc (occurs a (del1 a xs)) = occurs a xs"
- -- {* Wrong. Precondition needed.*}
+ \<comment> \<open>Wrong. Precondition needed.\<close>
quickcheck[random, expect = counterexample]
quickcheck[exhaustive, expect = counterexample]
oops
@@ -95,7 +95,7 @@
lemma "xs ~= [] \<longrightarrow> Suc (occurs a (del1 a xs)) = occurs a xs"
quickcheck[random, expect = counterexample]
quickcheck[exhaustive, expect = counterexample]
- -- {* Also wrong.*}
+ \<comment> \<open>Also wrong.\<close>
oops
lemma "0 < occurs a xs \<longrightarrow> Suc (occurs a (del1 a xs)) = occurs a xs"
@@ -111,7 +111,7 @@
lemma "occurs a xs = occurs b (replace a b xs)"
quickcheck[random, expect = counterexample]
quickcheck[exhaustive, expect = counterexample]
- -- {* Wrong. Precondition needed.*}
+ \<comment> \<open>Wrong. Precondition needed.\<close>
oops
lemma "occurs b xs = 0 \<or> a=b \<longrightarrow> occurs a xs = occurs b (replace a b xs)"
@@ -120,7 +120,7 @@
by (induct xs) simp_all
-subsection {* Trees *}
+subsection \<open>Trees\<close>
datatype 'a tree = Twig | Leaf 'a | Branch "'a tree" "'a tree"
@@ -141,13 +141,13 @@
theorem "plant (rev (leaves xt)) = mirror xt"
quickcheck[random, expect = counterexample]
quickcheck[exhaustive, expect = counterexample]
- --{* Wrong! *}
+ \<comment>\<open>Wrong!\<close>
oops
theorem "plant((leaves xt) @ (leaves yt)) = Branch xt yt"
quickcheck[random, expect = counterexample]
quickcheck[exhaustive, expect = counterexample]
- --{* Wrong! *}
+ \<comment>\<open>Wrong!\<close>
oops
datatype 'a ntree = Tip "'a" | Node "'a" "'a ntree" "'a ntree"
@@ -163,14 +163,14 @@
theorem "hd (inOrder xt) = root xt"
quickcheck[random, expect = counterexample]
quickcheck[exhaustive, expect = counterexample]
- --{* Wrong! *}
+ \<comment>\<open>Wrong!\<close>
oops
-subsection {* Exhaustive Testing beats Random Testing *}
+subsection \<open>Exhaustive Testing beats Random Testing\<close>
-text {* Here are some examples from mutants from the List theory
-where exhaustive testing beats random testing *}
+text \<open>Here are some examples from mutants from the List theory
+where exhaustive testing beats random testing\<close>
lemma
"[] ~= xs ==> hd xs = last (x # xs)"
@@ -245,7 +245,7 @@
quickcheck[exhaustive, expect = counterexample]
oops
-subsection {* Random Testing beats Exhaustive Testing *}
+subsection \<open>Random Testing beats Exhaustive Testing\<close>
lemma mult_inj_if_coprime_nat:
"inj_on f A \<Longrightarrow> inj_on g B
@@ -254,11 +254,11 @@
quickcheck[random]
oops
-subsection {* Examples with quantifiers *}
+subsection \<open>Examples with quantifiers\<close>
-text {*
+text \<open>
These examples show that we can handle quantifiers.
-*}
+\<close>
lemma "(\<exists>x. P x) \<longrightarrow> (\<forall>x. P x)"
quickcheck[random, expect = counterexample]
@@ -276,7 +276,7 @@
oops
-subsection {* Examples with sets *}
+subsection \<open>Examples with sets\<close>
lemma
"{} = A Un - A"
@@ -289,7 +289,7 @@
oops
-subsection {* Examples with relations *}
+subsection \<open>Examples with relations\<close>
lemma
"acyclic (R :: ('a * 'a) set) ==> acyclic S ==> acyclic (R Un S)"
@@ -323,7 +323,7 @@
oops
-subsection {* Examples with the descriptive operator *}
+subsection \<open>Examples with the descriptive operator\<close>
lemma
"(THE x. x = a) = b"
@@ -331,7 +331,7 @@
quickcheck[exhaustive, expect = counterexample]
oops
-subsection {* Examples with Multisets *}
+subsection \<open>Examples with Multisets\<close>
lemma
"X + Y = Y + (Z :: 'a multiset)"
@@ -351,11 +351,11 @@
quickcheck[exhaustive, expect = counterexample]
oops
-subsection {* Examples with numerical types *}
+subsection \<open>Examples with numerical types\<close>
-text {*
+text \<open>
Quickcheck supports the common types nat, int, rat and real.
-*}
+\<close>
lemma
"(x :: nat) > 0 ==> y > 0 ==> z > 0 ==> x * x + y * y \<noteq> z * z"
@@ -391,7 +391,7 @@
quickcheck[exhaustive, expect = counterexample]
oops
-subsubsection {* floor and ceiling functions *}
+subsubsection \<open>floor and ceiling functions\<close>
lemma "\<lfloor>x\<rfloor> + \<lfloor>y\<rfloor> = \<lfloor>x + y :: rat\<rfloor>"
quickcheck[expect = counterexample]
@@ -409,7 +409,7 @@
quickcheck[expect = counterexample]
oops
-subsection {* Examples with abstract types *}
+subsection \<open>Examples with abstract types\<close>
lemma
"Dlist.length (Dlist.remove x xs) = Dlist.length xs - 1"
@@ -423,7 +423,7 @@
quickcheck[random]
oops
-subsection {* Examples with Records *}
+subsection \<open>Examples with Records\<close>
record point =
xpos :: nat
@@ -446,7 +446,7 @@
quickcheck[random, expect = counterexample]
oops
-subsection {* Examples with locales *}
+subsection \<open>Examples with locales\<close>
locale Truth
@@ -499,7 +499,7 @@
oops
-subsection {* Examples with HOL quantifiers *}
+subsection \<open>Examples with HOL quantifiers\<close>
lemma
"\<forall> xs ys. xs = [] --> xs = ys"
@@ -516,7 +516,7 @@
quickcheck[exhaustive, expect = counterexample]
oops
-subsection {* Examples with underspecified/partial functions *}
+subsection \<open>Examples with underspecified/partial functions\<close>
lemma
"xs = [] ==> hd xs \<noteq> x"
@@ -538,7 +538,7 @@
quickcheck[random, report = true, expect = no_counterexample]
oops
-text {* with the simple testing scheme *}
+text \<open>with the simple testing scheme\<close>
setup Exhaustive_Generators.setup_exhaustive_datatype_interpretation
declare [[quickcheck_full_support = false]]
@@ -560,7 +560,7 @@
declare [[quickcheck_full_support = true]]
-subsection {* Equality Optimisation *}
+subsection \<open>Equality Optimisation\<close>
lemma
"f x = y ==> y = (0 :: nat)"
--- a/src/HOL/Quickcheck_Examples/Quickcheck_Interfaces.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Quickcheck_Examples/Quickcheck_Interfaces.thy Thu May 26 17:51:22 2016 +0200
@@ -2,51 +2,51 @@
imports Main
begin
-subsection {* Checking a single proposition (TODO) *}
+subsection \<open>Checking a single proposition (TODO)\<close>
-subsection {* Checking multiple propositions in one batch *}
+subsection \<open>Checking multiple propositions in one batch\<close>
-text {*
+text \<open>
First, this requires to setup special generators for all datatypes via the following command.
-*}
+\<close>
setup Exhaustive_Generators.setup_bounded_forall_datatype_interpretation
-text {*
+text \<open>
Now, the function Quickcheck.mk_batch_validator : Proof.context -> term list -> (int -> bool) list option
takes formulas of type bool with free variables, and returns a list of testing functions.
-*}
+\<close>
-ML {*
+ML \<open>
val SOME testers = Quickcheck.mk_batch_validator @{context}
[@{term "x = (1 :: nat)"}, @{term "x = (0 :: nat)"}, @{term "x <= (5 :: nat)"}, @{term "0 \<le> (x :: nat)"}]
-*}
+\<close>
-text {*
+text \<open>
It is up to the user with which strategy the conjectures should be tested.
For example, one could check all conjectures up to a given size, and check the different conjectures in sequence.
This is implemented by:
-*}
+\<close>
-ML {*
+ML \<open>
fun check_upto f i j = if i > j then true else f i andalso check_upto f (i + 1) j
-*}
+\<close>
-ML_val {*
+ML_val \<open>
map (fn test => check_upto test 0 1) testers
-*}
-ML_val {*
+\<close>
+ML_val \<open>
map (fn test => check_upto test 0 2) testers
-*}
-ML_val {*
+\<close>
+ML_val \<open>
map (fn test => check_upto test 0 3) testers
-*}
-ML_val {*
+\<close>
+ML_val \<open>
map (fn test => check_upto test 0 7) testers
-*}
+\<close>
-text {* Note that all conjectures must be executable to obtain the testers with the function above. *}
+text \<open>Note that all conjectures must be executable to obtain the testers with the function above.\<close>
end
--- a/src/HOL/Quickcheck_Examples/Quickcheck_Lattice_Examples.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Quickcheck_Examples/Quickcheck_Lattice_Examples.thy Thu May 26 17:51:22 2016 +0200
@@ -9,8 +9,8 @@
declare [[quickcheck_finite_type_size=5]]
-text {* We show how other default types help to find counterexamples to propositions if
- the standard default type @{typ int} is insufficient. *}
+text \<open>We show how other default types help to find counterexamples to propositions if
+ the standard default type @{typ int} is insufficient.\<close>
notation
less_eq (infix "\<sqsubseteq>" 50) and
@@ -22,7 +22,7 @@
declare [[quickcheck_narrowing_active = false, quickcheck_timeout = 3600]]
-subsection {* Distributive lattices *}
+subsection \<open>Distributive lattices\<close>
lemma sup_inf_distrib2:
"((y :: 'a :: distrib_lattice) \<sqinter> z) \<squnion> x = (y \<squnion> x) \<sqinter> (z \<squnion> x)"
@@ -49,7 +49,7 @@
quickcheck[expect = counterexample]
oops
-subsection {* Bounded lattices *}
+subsection \<open>Bounded lattices\<close>
lemma inf_bot_left [simp]:
"\<bottom> \<sqinter> (x :: 'a :: bounded_lattice_bot) = \<bottom>"
--- a/src/HOL/Quickcheck_Examples/Quickcheck_Narrowing_Examples.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Quickcheck_Examples/Quickcheck_Narrowing_Examples.thy Thu May 26 17:51:22 2016 +0200
@@ -3,7 +3,7 @@
Copyright 2011 TU Muenchen
*)
-section {* Examples for narrowing-based testing *}
+section \<open>Examples for narrowing-based testing\<close>
theory Quickcheck_Narrowing_Examples
imports Main
@@ -11,7 +11,7 @@
declare [[quickcheck_timeout = 3600]]
-subsection {* Minimalistic examples *}
+subsection \<open>Minimalistic examples\<close>
lemma
"x = y"
@@ -23,7 +23,7 @@
quickcheck[tester = narrowing, finite_types = false, default_type = nat, expect = counterexample]
oops
-subsection {* Simple examples with existentials *}
+subsection \<open>Simple examples with existentials\<close>
lemma
"\<exists> y :: nat. \<forall> x. x = y"
@@ -50,19 +50,19 @@
quickcheck[tester = narrowing, finite_types = false, size = 7]
oops
-text {* A false conjecture derived from an partial copy-n-paste of @{thm not_distinct_decomp} *}
+text \<open>A false conjecture derived from an partial copy-n-paste of @{thm not_distinct_decomp}\<close>
lemma
"~ distinct ws ==> EX xs ys zs y. ws = xs @ [y] @ ys @ [y]"
quickcheck[tester = narrowing, finite_types = false, default_type = nat, expect = counterexample]
oops
-text {* A false conjecture derived from theorems @{thm split_list_first} and @{thm split_list_last} *}
+text \<open>A false conjecture derived from theorems @{thm split_list_first} and @{thm split_list_last}\<close>
lemma
"x : set xs ==> EX ys zs. xs = ys @ x # zs & x ~: set zs & x ~: set ys"
quickcheck[tester = narrowing, finite_types = false, default_type = nat, expect = counterexample]
oops
-text {* A false conjecture derived from @{thm map_eq_Cons_conv} with a typo *}
+text \<open>A false conjecture derived from @{thm map_eq_Cons_conv} with a typo\<close>
lemma
"(map f xs = y # ys) = (EX z zs. xs = z' # zs & f z = y & map f zs = ys)"
quickcheck[tester = narrowing, finite_types = false, default_type = nat, expect = counterexample]
@@ -73,7 +73,7 @@
quickcheck[exhaustive, random, expect = no_counterexample]
oops
-subsection {* Simple list examples *}
+subsection \<open>Simple list examples\<close>
lemma "rev xs = xs"
quickcheck[tester = narrowing, finite_types = false, default_type = nat, expect = counterexample]
@@ -87,7 +87,7 @@
quickcheck[tester = narrowing, finite_types = true, expect = counterexample]
oops
-subsection {* Simple examples with functions *}
+subsection \<open>Simple examples with functions\<close>
lemma "map f xs = map g xs"
quickcheck[tester = narrowing, finite_types = false, expect = counterexample]
@@ -117,7 +117,7 @@
oops
*)
-subsection {* Simple examples with inductive predicates *}
+subsection \<open>Simple examples with inductive predicates\<close>
inductive even where
"even 0" |
@@ -130,7 +130,7 @@
oops
-subsection {* AVL Trees *}
+subsection \<open>AVL Trees\<close>
datatype 'a tree = ET | MKT 'a "'a tree" "'a tree" nat
@@ -204,7 +204,7 @@
else MKT n l r' (1 + max hl hr'))"
-subsubsection {* Necessary setup for code generation *}
+subsubsection \<open>Necessary setup for code generation\<close>
primrec set_of'
where
@@ -221,14 +221,14 @@
declare is_ord.simps(1)[code] is_ord_mkt[code]
-subsubsection {* Invalid Lemma due to typo in @{const l_bal} *}
+subsubsection \<open>Invalid Lemma due to typo in @{const l_bal}\<close>
lemma is_ord_l_bal:
"\<lbrakk> is_ord(MKT (x :: nat) l r h); height l = height r + 2 \<rbrakk> \<Longrightarrow> is_ord(l_bal(x,l,r))"
quickcheck[tester = narrowing, finite_types = false, default_type = nat, size = 6, expect = counterexample]
oops
-subsection {* Examples with hd and last *}
+subsection \<open>Examples with hd and last\<close>
lemma
"hd (xs @ ys) = hd ys"
@@ -240,7 +240,7 @@
quickcheck[narrowing, expect = counterexample]
oops
-subsection {* Examples with underspecified/partial functions *}
+subsection \<open>Examples with underspecified/partial functions\<close>
lemma
"xs = [] ==> hd xs \<noteq> x"
--- a/src/HOL/Quotient_Examples/DList.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Quotient_Examples/DList.thy Thu May 26 17:51:22 2016 +0200
@@ -5,13 +5,13 @@
and theory morphism version by Maksym Bortin
*)
-section {* Lists with distinct elements *}
+section \<open>Lists with distinct elements\<close>
theory DList
imports "~~/src/HOL/Library/Quotient_List"
begin
-text {* Some facts about lists *}
+text \<open>Some facts about lists\<close>
lemma remdups_removeAll_commute[simp]:
"remdups (removeAll x l) = removeAll x (remdups l)"
@@ -66,7 +66,7 @@
by (induct xa ya rule: list_induct2')
(metis remdups_nil_noteq_cons set_remdups)+
-text {* Setting up the quotient type *}
+text \<open>Setting up the quotient type\<close>
definition
dlist_eq :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool"
@@ -93,7 +93,7 @@
'a dlist = "'a list" / "dlist_eq"
by (rule dlist_eq_equivp)
-text {* respectfulness and constant definitions *}
+text \<open>respectfulness and constant definitions\<close>
definition [simp]: "card_remdups = length \<circ> remdups"
definition [simp]: "foldr_remdups f xs e = foldr f (remdups xs) e"
@@ -125,7 +125,7 @@
quotient_definition "list_of_dlist :: 'a dlist \<Rightarrow> 'a list"
is "remdups" by simp
-text {* lifted theorems *}
+text \<open>lifted theorems\<close>
lemma dlist_member_insert:
"member dl x \<Longrightarrow> insert x dl = dl"
--- a/src/HOL/Quotient_Examples/FSet.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Quotient_Examples/FSet.thy Thu May 26 17:51:22 2016 +0200
@@ -16,10 +16,10 @@
imports "~~/src/HOL/Library/Multiset" "~~/src/HOL/Library/Quotient_List"
begin
-text {*
+text \<open>
The type of finite sets is created by a quotient construction
over lists. The definition of the equivalence:
-*}
+\<close>
definition
list_eq :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" (infix "\<approx>" 50)
@@ -42,17 +42,17 @@
"equivp list_eq"
by (auto intro: equivpI list_eq_reflp list_eq_symp list_eq_transp)
-text {* The @{text fset} type *}
+text \<open>The \<open>fset\<close> type\<close>
quotient_type
'a fset = "'a list" / "list_eq"
by (rule list_eq_equivp)
-text {*
+text \<open>
Definitions for sublist, cardinality,
intersection, difference and respectful fold over
lists.
-*}
+\<close>
declare List.member_def [simp]
@@ -104,7 +104,7 @@
by (simp add: fold_once_def)
-section {* Quotient composition lemmas *}
+section \<open>Quotient composition lemmas\<close>
lemma list_all2_refl':
assumes q: "equivp R"
@@ -186,10 +186,10 @@
by (rule quotient_compose_list_g, rule Quotient3_fset, rule list_eq_equivp)
-section {* Quotient definitions for fsets *}
+section \<open>Quotient definitions for fsets\<close>
-subsection {* Finite sets are a bounded, distributive lattice with minus *}
+subsection \<open>Finite sets are a bounded, distributive lattice with minus\<close>
instantiation fset :: (type) "{bounded_lattice_bot, distrib_lattice, minus}"
begin
@@ -282,7 +282,7 @@
end
-subsection {* Other constants for fsets *}
+subsection \<open>Other constants for fsets\<close>
quotient_definition
"insert_fset :: 'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
@@ -311,7 +311,7 @@
"x |\<notin>| S \<equiv> \<not> (x |\<in>| S)"
-subsection {* Other constants on the Quotient Type *}
+subsection \<open>Other constants on the Quotient Type\<close>
quotient_definition
"card_fset :: 'a fset \<Rightarrow> nat"
@@ -395,7 +395,7 @@
is filter by force
-subsection {* Compositional respectfulness and preservation lemmas *}
+subsection \<open>Compositional respectfulness and preservation lemmas\<close>
lemma Nil_rsp2 [quot_respect]:
shows "(list_all2 op \<approx> OOO op \<approx>) Nil Nil"
@@ -494,9 +494,9 @@
by (auto simp add: fun_eq_iff)
(simp only: map_map[symmetric] map_prs_aux[OF Quotient3_fset Quotient3_fset])
-section {* Lifted theorems *}
+section \<open>Lifted theorems\<close>
-subsection {* fset *}
+subsection \<open>fset\<close>
lemma fset_simps [simp]:
shows "fset {||} = {}"
@@ -532,7 +532,7 @@
by (descending) (auto)
-subsection {* in_fset *}
+subsection \<open>in_fset\<close>
lemma in_fset:
shows "x |\<in>| S \<longleftrightarrow> x \<in> fset S"
@@ -555,7 +555,7 @@
by descending simp
-subsection {* insert_fset *}
+subsection \<open>insert_fset\<close>
lemma in_insert_fset_iff [simp]:
shows "x |\<in>| insert_fset y S \<longleftrightarrow> x = y \<or> x |\<in>| S"
@@ -592,7 +592,7 @@
by (descending) (auto)
-subsection {* union_fset *}
+subsection \<open>union_fset\<close>
lemmas [simp] =
sup_bot_left[where 'a="'a fset"]
@@ -615,7 +615,7 @@
by (descending) (simp)
-subsection {* minus_fset *}
+subsection \<open>minus_fset\<close>
lemma minus_in_fset:
shows "x |\<in>| (xs - ys) \<longleftrightarrow> x |\<in>| xs \<and> x |\<notin>| ys"
@@ -644,7 +644,7 @@
by blast
-subsection {* remove_fset *}
+subsection \<open>remove_fset\<close>
lemma in_remove_fset:
shows "x |\<in>| remove_fset y S \<longleftrightarrow> x |\<in>| S \<and> x \<noteq> y"
@@ -663,7 +663,7 @@
by (descending) (auto simp add: insert_absorb)
-subsection {* inter_fset *}
+subsection \<open>inter_fset\<close>
lemma inter_empty_fset_l:
shows "{||} |\<inter>| S = {||}"
@@ -682,7 +682,7 @@
by (descending) (simp)
-subsection {* subset_fset and psubset_fset *}
+subsection \<open>subset_fset and psubset_fset\<close>
lemma subset_fset:
shows "xs |\<subseteq>| ys \<longleftrightarrow> fset xs \<subseteq> fset ys"
@@ -710,7 +710,7 @@
by (metis fset_simps(1) psubset_fset not_psubset_empty)
-subsection {* map_fset *}
+subsection \<open>map_fset\<close>
lemma map_fset_simps [simp]:
shows "map_fset f {||} = {||}"
@@ -733,7 +733,7 @@
by descending auto
-subsection {* card_fset *}
+subsection \<open>card_fset\<close>
lemma card_fset:
shows "card_fset xs = card (fset xs)"
@@ -852,7 +852,7 @@
by (rule card_Diff_subset_Int) (simp)
-subsection {* concat_fset *}
+subsection \<open>concat_fset\<close>
lemma concat_empty_fset [simp]:
shows "concat_fset {||} = {||}"
@@ -870,7 +870,7 @@
shows "map_fset f (concat_fset xs) = concat_fset (map_fset (map_fset f) xs)"
by (lifting map_concat)
-subsection {* filter_fset *}
+subsection \<open>filter_fset\<close>
lemma subset_filter_fset:
"filter_fset P xs |\<subseteq>| filter_fset Q xs = (\<forall> x. x |\<in>| xs \<longrightarrow> P x \<longrightarrow> Q x)"
@@ -886,7 +886,7 @@
unfolding less_fset_def by (auto simp add: subset_filter_fset eq_filter_fset)
-subsection {* fold_fset *}
+subsection \<open>fold_fset\<close>
lemma fold_empty_fset:
"fold_fset f {||} = id"
@@ -907,7 +907,7 @@
proof -
from assms have "fold f (remdups xs) = fold f (remove1 x (remdups xs)) \<circ> f x"
by (auto intro!: fold_remove1_split elim: rsp_foldE)
- then show ?thesis using `rsp_fold f` by (simp add: fold_once_fold_remdups remdups_removeAll)
+ then show ?thesis using \<open>rsp_fold f\<close> by (simp add: fold_once_fold_remdups remdups_removeAll)
qed
lemma in_commute_fold_fset:
@@ -915,7 +915,7 @@
by descending (simp add: member_commute_fold_once)
-subsection {* Choice in fsets *}
+subsection \<open>Choice in fsets\<close>
lemma fset_choice:
assumes a: "\<forall>x. x |\<in>| A \<longrightarrow> (\<exists>y. P x y)"
@@ -926,7 +926,7 @@
by (auto simp add: Ball_def)
-section {* Induction and Cases rules for fsets *}
+section \<open>Induction and Cases rules for fsets\<close>
lemma fset_exhaust [case_names empty insert, cases type: fset]:
assumes empty_fset_case: "S = {||} \<Longrightarrow> P"
@@ -1028,7 +1028,7 @@
apply simp_all
done
-text {* Extensionality *}
+text \<open>Extensionality\<close>
lemma fset_eqI:
assumes "\<And>x. x \<in> fset A \<longleftrightarrow> x \<in> fset B"
@@ -1045,8 +1045,8 @@
with A show ?case by (metis in_fset_mdef)
qed
-subsection {* alternate formulation with a different decomposition principle
- and a proof of equivalence *}
+subsection \<open>alternate formulation with a different decomposition principle
+ and a proof of equivalence\<close>
inductive
list_eq2 :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" ("_ \<approx>2 _")
@@ -1098,7 +1098,7 @@
have "card_list l = 0" by fact
then have "\<forall>x. \<not> List.member l x" by auto
then have z: "l = []" by auto
- then have "r = []" using `l \<approx> r` by simp
+ then have "r = []" using \<open>l \<approx> r\<close> by simp
then show ?case using z list_eq2_refl by simp
next
case (Suc m)
@@ -1151,10 +1151,10 @@
using assms
by (lifting list_eq2.induct[simplified list_eq2_equiv[symmetric]])
-ML {*
+ML \<open>
fun dest_fsetT (Type (@{type_name fset}, [T])) = T
| dest_fsetT T = raise TYPE ("dest_fsetT: fset type expected", [T], []);
-*}
+\<close>
no_notation
list_eq (infix "\<approx>" 50) and
--- a/src/HOL/Quotient_Examples/Int_Pow.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Quotient_Examples/Int_Pow.thy Thu May 26 17:51:22 2016 +0200
@@ -51,7 +51,7 @@
assumes [simp, intro]: "a \<in> Units G"
shows "a (^) (m::nat) \<otimes> inv (a (^) (n::nat)) = inv (a (^) n) \<otimes> a (^) m"
proof (cases "m\<ge>n")
- have [simp]: "a \<in> carrier G" using `a \<in> _` by (rule Units_closed)
+ have [simp]: "a \<in> carrier G" using \<open>a \<in> _\<close> by (rule Units_closed)
case True
then obtain k where *:"m = k + n" and **:"m = n + k" by (metis le_iff_add add.commute)
have "a (^) (m::nat) \<otimes> inv (a (^) (n::nat)) = a (^) k"
@@ -60,7 +60,7 @@
using ** by (auto simp add: nat_pow_mult[symmetric] m_assoc[symmetric])
finally show ?thesis .
next
- have [simp]: "a \<in> carrier G" using `a \<in> _` by (rule Units_closed)
+ have [simp]: "a \<in> carrier G" using \<open>a \<in> _\<close> by (rule Units_closed)
case False
then obtain k where *:"n = k + m" and **:"n = m + k"
by (metis le_iff_add add.commute nat_le_linear)
@@ -84,23 +84,23 @@
assumes a_in_G [simp, intro]: "a \<in> Units G"
shows "a (^) b \<otimes> inv (a (^) c) = a (^) d \<otimes> inv (a (^) e)"
proof(cases "b\<ge>c")
- have [simp]: "a \<in> carrier G" using `a \<in> _` by (rule Units_closed)
+ have [simp]: "a \<in> carrier G" using \<open>a \<in> _\<close> by (rule Units_closed)
case True
then obtain n where "b = n + c" by (metis le_iff_add add.commute)
then have "d = n + e" using eq by arith
- from `b = _` have "a (^) b \<otimes> inv (a (^) c) = a (^) n"
+ from \<open>b = _\<close> have "a (^) b \<otimes> inv (a (^) c) = a (^) n"
by (auto simp add: nat_pow_mult[symmetric] m_assoc)
- also from `d = _` have "\<dots> = a (^) d \<otimes> inv (a (^) e)"
+ also from \<open>d = _\<close> have "\<dots> = a (^) d \<otimes> inv (a (^) e)"
by (auto simp add: nat_pow_mult[symmetric] m_assoc)
finally show ?thesis .
next
- have [simp]: "a \<in> carrier G" using `a \<in> _` by (rule Units_closed)
+ have [simp]: "a \<in> carrier G" using \<open>a \<in> _\<close> by (rule Units_closed)
case False
then obtain n where "c = n + b" by (metis le_iff_add add.commute nat_le_linear)
then have "e = n + d" using eq by arith
- from `c = _` have "a (^) b \<otimes> inv (a (^) c) = inv (a (^) n)"
+ from \<open>c = _\<close> have "a (^) b \<otimes> inv (a (^) c) = inv (a (^) n)"
by (auto simp add: nat_pow_mult[symmetric] m_assoc[symmetric] inv_mult_units)
- also from `e = _` have "\<dots> = a (^) d \<otimes> inv (a (^) e)"
+ also from \<open>e = _\<close> have "\<dots> = a (^) d \<otimes> inv (a (^) e)"
by (auto simp add: nat_pow_mult[symmetric] m_assoc[symmetric] inv_mult_units)
finally show ?thesis .
qed
--- a/src/HOL/Quotient_Examples/Lift_DList.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Quotient_Examples/Lift_DList.thy Thu May 26 17:51:22 2016 +0200
@@ -6,7 +6,7 @@
imports Main
begin
-subsection {* The type of distinct lists *}
+subsection \<open>The type of distinct lists\<close>
typedef 'a dlist = "{xs::'a list. distinct xs}"
morphisms list_of_dlist Abs_dlist
@@ -16,7 +16,7 @@
setup_lifting type_definition_dlist
-text {* Fundamental operations: *}
+text \<open>Fundamental operations:\<close>
lift_definition empty :: "'a dlist" is "[]"
by simp
@@ -33,7 +33,7 @@
lift_definition filter :: "('a \<Rightarrow> bool) \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" is List.filter
by simp
-text {* Derived operations: *}
+text \<open>Derived operations:\<close>
lift_definition null :: "'a dlist \<Rightarrow> bool" is List.null .
@@ -47,7 +47,7 @@
lift_definition concat :: "'a dlist dlist \<Rightarrow> 'a dlist" is "remdups o List.concat" by auto
-text {* We can export code: *}
+text \<open>We can export code:\<close>
export_code empty insert remove map filter null member length fold foldr concat in SML
--- a/src/HOL/Quotient_Examples/Lift_FSet.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Quotient_Examples/Lift_FSet.thy Thu May 26 17:51:22 2016 +0200
@@ -2,13 +2,13 @@
Author: Brian Huffman, TU Munich
*)
-section {* Lifting and transfer with a finite set type *}
+section \<open>Lifting and transfer with a finite set type\<close>
theory Lift_FSet
imports Main
begin
-subsection {* Equivalence relation and quotient type definition *}
+subsection \<open>Equivalence relation and quotient type definition\<close>
definition list_eq :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool"
where [simp]: "list_eq xs ys \<longleftrightarrow> set xs = set ys"
@@ -37,7 +37,7 @@
quotient_type 'a fset = "'a list" / "list_eq" parametric list_eq_transfer
by (rule equivp_list_eq)
-subsection {* Lifted constant definitions *}
+subsection \<open>Lifted constant definitions\<close>
lift_definition fnil :: "'a fset" ("{||}") is "[]" parametric list.ctr_transfer(1) .
@@ -56,8 +56,8 @@
lift_definition fset :: "'a fset \<Rightarrow> 'a set" is set parametric list.set_transfer
by simp
-text {* Constants with nested types (like concat) yield a more
- complicated proof obligation. *}
+text \<open>Constants with nested types (like concat) yield a more
+ complicated proof obligation.\<close>
lemma list_all2_cr_fset:
"list_all2 cr_fset xs ys \<longleftrightarrow> map abs_fset xs = ys"
@@ -112,34 +112,32 @@
lemma fmember_fmap[simp]: "a |\<in>| fmap f X = (\<exists>b. b |\<in>| X \<and> a = f b)"
by transfer auto
-text {* We can export code: *}
+text \<open>We can export code:\<close>
export_code fnil fcons fappend fmap ffilter fset fmember in SML
-subsection {* Using transfer with type @{text "fset"} *}
+subsection \<open>Using transfer with type \<open>fset\<close>\<close>
-text {* The correspondence relation @{text "cr_fset"} can only relate
- @{text "list"} and @{text "fset"} types with the same element type.
- To relate nested types like @{text "'a list list"} and
- @{text "'a fset fset"}, we define a parameterized version of the
- correspondence relation, @{text "pcr_fset"}. *}
+text \<open>The correspondence relation \<open>cr_fset\<close> can only relate
+ \<open>list\<close> and \<open>fset\<close> types with the same element type.
+ To relate nested types like \<open>'a list list\<close> and
+ \<open>'a fset fset\<close>, we define a parameterized version of the
+ correspondence relation, \<open>pcr_fset\<close>.\<close>
thm pcr_fset_def
-subsection {* Transfer examples *}
+subsection \<open>Transfer examples\<close>
-text {* The @{text "transfer"} method replaces equality on @{text
- "fset"} with the @{text "list_eq"} relation on lists, which is
- logically equivalent. *}
+text \<open>The \<open>transfer\<close> method replaces equality on \<open>fset\<close> with the \<open>list_eq\<close> relation on lists, which is
+ logically equivalent.\<close>
lemma "fmap f (fmap g xs) = fmap (f \<circ> g) xs"
apply transfer
apply simp
done
-text {* The @{text "transfer'"} variant can replace equality on @{text
- "fset"} with equality on @{text "list"}, which is logically stronger
- but sometimes more convenient. *}
+text \<open>The \<open>transfer'\<close> variant can replace equality on \<open>fset\<close> with equality on \<open>list\<close>, which is logically stronger
+ but sometimes more convenient.\<close>
lemma "fmap f (fmap g xs) = fmap (f \<circ> g) xs"
using map_map [Transfer.transferred] .
--- a/src/HOL/Quotient_Examples/Lift_Fun.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Quotient_Examples/Lift_Fun.thy Thu May 26 17:51:22 2016 +0200
@@ -2,21 +2,21 @@
Author: Ondrej Kuncar
*)
-section {* Example of lifting definitions with contravariant or co/contravariant type variables *}
+section \<open>Example of lifting definitions with contravariant or co/contravariant type variables\<close>
theory Lift_Fun
imports Main "~~/src/HOL/Library/Quotient_Syntax"
begin
-text {* This file is meant as a test case.
+text \<open>This file is meant as a test case.
It contains examples of lifting definitions with quotients that have contravariant
- type variables or type variables which are covariant and contravariant in the same time. *}
+ type variables or type variables which are covariant and contravariant in the same time.\<close>
-subsection {* Contravariant type variables *}
+subsection \<open>Contravariant type variables\<close>
-text {* 'a is a contravariant type variable and we are able to map over this variable
- in the following four definitions. This example is based on HOL/Fun.thy. *}
+text \<open>'a is a contravariant type variable and we are able to map over this variable
+ in the following four definitions. This example is based on HOL/Fun.thy.\<close>
quotient_type
('a, 'b) fun' (infixr "\<rightarrow>" 55) = "'a \<Rightarrow> 'b" / "op ="
@@ -36,10 +36,10 @@
quotient_definition "bij_betw' :: ('a \<rightarrow> 'b) \<rightarrow> 'a set \<rightarrow> 'b set \<rightarrow> bool" is bij_betw done
-subsection {* Co/Contravariant type variables *}
+subsection \<open>Co/Contravariant type variables\<close>
-text {* 'a is a covariant and contravariant type variable in the same time.
- The following example is a bit artificial. We haven't had a natural one yet. *}
+text \<open>'a is a covariant and contravariant type variable in the same time.
+ The following example is a bit artificial. We haven't had a natural one yet.\<close>
quotient_type 'a endofun = "'a \<Rightarrow> 'a" / "op =" by (simp add: identity_equivp)
@@ -49,7 +49,7 @@
quotient_definition "map_endofun :: ('a \<Rightarrow> 'b) \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a endofun \<Rightarrow> 'b endofun" is
map_endofun' done
-text {* Registration of the map function for 'a endofun. *}
+text \<open>Registration of the map function for 'a endofun.\<close>
functor map_endofun : map_endofun
proof -
@@ -63,7 +63,7 @@
by (auto simp: map_endofun_def map_endofun'_def map_fun_def fun_eq_iff) (simp add: a o_assoc)
qed
-text {* Relator for 'a endofun. *}
+text \<open>Relator for 'a endofun.\<close>
definition
rel_endofun' :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> bool"
@@ -110,8 +110,8 @@
quotient_type 'a endofun' = "'a endofun" / "op =" by (simp add: identity_equivp)
-text {* We have to map "'a endofun" to "('a endofun') endofun", i.e., mapping (lifting)
- over a type variable which is a covariant and contravariant type variable. *}
+text \<open>We have to map "'a endofun" to "('a endofun') endofun", i.e., mapping (lifting)
+ over a type variable which is a covariant and contravariant type variable.\<close>
quotient_definition "endofun'_id_id :: ('a endofun') endofun'" is endofun_id_id done
--- a/src/HOL/Quotient_Examples/Lift_Set.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Quotient_Examples/Lift_Set.thy Thu May 26 17:51:22 2016 +0200
@@ -2,7 +2,7 @@
Author: Lukas Bulwahn and Ondrej Kuncar
*)
-section {* Example of lifting definitions with the lifting infrastructure *}
+section \<open>Example of lifting definitions with the lifting infrastructure\<close>
theory Lift_Set
imports Main
@@ -16,7 +16,7 @@
setup_lifting type_definition_set[unfolded set_def]
-text {* Now, we can employ lift_definition to lift definitions. *}
+text \<open>Now, we can employ lift_definition to lift definitions.\<close>
lift_definition empty :: "'a set" is "bot :: 'a \<Rightarrow> bool" .
--- a/src/HOL/Quotient_Examples/Quotient_Int.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Quotient_Examples/Quotient_Int.thy Thu May 26 17:51:22 2016 +0200
@@ -108,7 +108,7 @@
end
-text{* The integers form a @{text comm_ring_1}*}
+text\<open>The integers form a \<open>comm_ring_1\<close>\<close>
instance int :: comm_ring_1
proof
@@ -235,7 +235,7 @@
using a
by (drule_tac zero_le_imp_eq_int) (auto simp add: zmult_zless_mono2_lemma)
-text{*The integers form an ordered integral domain*}
+text\<open>The integers form an ordered integral domain\<close>
instance int :: linordered_idom
proof
@@ -276,7 +276,7 @@
by (descending) (auto intro: int_induct2)
-text {* Magnitide of an Integer, as a Natural Number: @{term nat} *}
+text \<open>Magnitide of an Integer, as a Natural Number: @{term nat}\<close>
definition
"int_to_nat_raw \<equiv> \<lambda>(x, y).x - (y::nat)"
--- a/src/HOL/Quotient_Examples/Quotient_Message.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Quotient_Examples/Quotient_Message.thy Thu May 26 17:51:22 2016 +0200
@@ -8,7 +8,7 @@
imports Main "~~/src/HOL/Library/Quotient_Syntax"
begin
-subsection{*Defining the Free Algebra*}
+subsection\<open>Defining the Free Algebra\<close>
datatype
freemsg = NONCE nat
@@ -30,7 +30,7 @@
lemmas msgrel.intros[intro]
-text{*Proving that it is an equivalence relation*}
+text\<open>Proving that it is an equivalence relation\<close>
lemma msgrel_refl: "X \<sim> X"
by (induct X) (auto intro: msgrel.intros)
@@ -42,9 +42,9 @@
show "transp msgrel" by (rule transpI) (blast intro: msgrel.TRANS)
qed
-subsection{*Some Functions on the Free Algebra*}
+subsection\<open>Some Functions on the Free Algebra\<close>
-subsubsection{*The Set of Nonces*}
+subsubsection\<open>The Set of Nonces\<close>
primrec
freenonces :: "freemsg \<Rightarrow> nat set"
@@ -59,10 +59,10 @@
shows "freenonces U = freenonces V"
using a by (induct) (auto)
-subsubsection{*The Left Projection*}
+subsubsection\<open>The Left Projection\<close>
-text{*A function to return the left part of the top pair in a message. It will
-be lifted to the initial algebra, to serve as an example of that process.*}
+text\<open>A function to return the left part of the top pair in a message. It will
+be lifted to the initial algebra, to serve as an example of that process.\<close>
primrec
freeleft :: "freemsg \<Rightarrow> freemsg"
where
@@ -71,9 +71,9 @@
| "freeleft (CRYPT K X) = freeleft X"
| "freeleft (DECRYPT K X) = freeleft X"
-text{*This theorem lets us prove that the left function respects the
+text\<open>This theorem lets us prove that the left function respects the
equivalence relation. It also helps us prove that MPair
- (the abstract constructor) is injective*}
+ (the abstract constructor) is injective\<close>
lemma msgrel_imp_eqv_freeleft_aux:
shows "freeleft U \<sim> freeleft U"
by (fact msgrel_refl)
@@ -84,9 +84,9 @@
using a
by (induct) (auto intro: msgrel_imp_eqv_freeleft_aux)
-subsubsection{*The Right Projection*}
+subsubsection\<open>The Right Projection\<close>
-text{*A function to return the right part of the top pair in a message.*}
+text\<open>A function to return the right part of the top pair in a message.\<close>
primrec
freeright :: "freemsg \<Rightarrow> freemsg"
where
@@ -95,9 +95,9 @@
| "freeright (CRYPT K X) = freeright X"
| "freeright (DECRYPT K X) = freeright X"
-text{*This theorem lets us prove that the right function respects the
+text\<open>This theorem lets us prove that the right function respects the
equivalence relation. It also helps us prove that MPair
- (the abstract constructor) is injective*}
+ (the abstract constructor) is injective\<close>
lemma msgrel_imp_eqv_freeright_aux:
shows "freeright U \<sim> freeright U"
by (fact msgrel_refl)
@@ -108,9 +108,9 @@
using a
by (induct) (auto intro: msgrel_imp_eqv_freeright_aux)
-subsubsection{*The Discriminator for Constructors*}
+subsubsection\<open>The Discriminator for Constructors\<close>
-text{*A function to distinguish nonces, mpairs and encryptions*}
+text\<open>A function to distinguish nonces, mpairs and encryptions\<close>
primrec
freediscrim :: "freemsg \<Rightarrow> int"
where
@@ -119,18 +119,18 @@
| "freediscrim (CRYPT K X) = freediscrim X + 2"
| "freediscrim (DECRYPT K X) = freediscrim X - 2"
-text{*This theorem helps us prove @{term "Nonce N \<noteq> MPair X Y"}*}
+text\<open>This theorem helps us prove @{term "Nonce N \<noteq> MPair X Y"}\<close>
theorem msgrel_imp_eq_freediscrim:
assumes a: "U \<sim> V"
shows "freediscrim U = freediscrim V"
using a by (induct) (auto)
-subsection{*The Initial Algebra: A Quotiented Message Type*}
+subsection\<open>The Initial Algebra: A Quotiented Message Type\<close>
quotient_type msg = freemsg / msgrel
by (rule equiv_msgrel)
-text{*The abstract message constructors*}
+text\<open>The abstract message constructors\<close>
quotient_definition
"Nonce :: nat \<Rightarrow> msg"
@@ -156,7 +156,7 @@
"DECRYPT"
by (rule DECRYPT)
-text{*Establishing these two equations is the point of the whole exercise*}
+text\<open>Establishing these two equations is the point of the whole exercise\<close>
theorem CD_eq [simp]:
shows "Crypt K (Decrypt K X) = X"
by (lifting CD)
@@ -165,7 +165,7 @@
shows "Decrypt K (Crypt K X) = X"
by (lifting DC)
-subsection{*The Abstract Function to Return the Set of Nonces*}
+subsection\<open>The Abstract Function to Return the Set of Nonces\<close>
quotient_definition
"nonces:: msg \<Rightarrow> nat set"
@@ -173,7 +173,7 @@
"freenonces"
by (rule msgrel_imp_eq_freenonces)
-text{*Now prove the four equations for @{term nonces}*}
+text\<open>Now prove the four equations for @{term nonces}\<close>
lemma nonces_Nonce [simp]:
shows "nonces (Nonce N) = {N}"
@@ -191,7 +191,7 @@
shows "nonces (Decrypt K X) = nonces X"
by (lifting freenonces.simps(4))
-subsection{*The Abstract Function to Return the Left Part*}
+subsection\<open>The Abstract Function to Return the Left Part\<close>
quotient_definition
"left:: msg \<Rightarrow> msg"
@@ -215,7 +215,7 @@
shows "left (Decrypt K X) = left X"
by (lifting freeleft.simps(4))
-subsection{*The Abstract Function to Return the Right Part*}
+subsection\<open>The Abstract Function to Return the Right Part\<close>
quotient_definition
"right:: msg \<Rightarrow> msg"
@@ -223,7 +223,7 @@
"freeright"
by (rule msgrel_imp_eqv_freeright)
-text{*Now prove the four equations for @{term right}*}
+text\<open>Now prove the four equations for @{term right}\<close>
lemma right_Nonce [simp]:
shows "right (Nonce N) = Nonce N"
@@ -241,9 +241,9 @@
shows "right (Decrypt K X) = right X"
by (lifting freeright.simps(4))
-subsection{*Injectivity Properties of Some Constructors*}
+subsection\<open>Injectivity Properties of Some Constructors\<close>
-text{*Can also be proved using the function @{term nonces}*}
+text\<open>Can also be proved using the function @{term nonces}\<close>
lemma Nonce_Nonce_eq [iff]:
shows "(Nonce m = Nonce n) = (m = n)"
proof
@@ -273,13 +273,13 @@
shows "Nonce N \<noteq> MPair X Y"
by (descending) (auto dest: msgrel_imp_eq_freediscrim)
-text{*Example suggested by a referee*}
+text\<open>Example suggested by a referee\<close>
theorem Crypt_Nonce_neq_Nonce:
shows "Crypt K (Nonce M) \<noteq> Nonce N"
by (descending) (auto dest: msgrel_imp_eq_freediscrim)
-text{*...and many similar results*}
+text\<open>...and many similar results\<close>
theorem Crypt2_Nonce_neq_Nonce:
shows "Crypt K (Crypt K' (Nonce M)) \<noteq> Nonce N"
@@ -317,10 +317,10 @@
by (descending) (auto intro: freemsg.induct)
-subsection{*The Abstract Discriminator*}
+subsection\<open>The Abstract Discriminator\<close>
-text{*However, as @{text Crypt_Nonce_neq_Nonce} above illustrates, we don't
-need this function in order to prove discrimination theorems.*}
+text\<open>However, as \<open>Crypt_Nonce_neq_Nonce\<close> above illustrates, we don't
+need this function in order to prove discrimination theorems.\<close>
quotient_definition
"discrim:: msg \<Rightarrow> int"
@@ -328,7 +328,7 @@
"freediscrim"
by (rule msgrel_imp_eq_freediscrim)
-text{*Now prove the four equations for @{term discrim}*}
+text\<open>Now prove the four equations for @{term discrim}\<close>
lemma discrim_Nonce [simp]:
shows "discrim (Nonce N) = 0"
--- a/src/HOL/SET_Protocol/Cardholder_Registration.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/SET_Protocol/Cardholder_Registration.thy Thu May 26 17:51:22 2016 +0200
@@ -5,33 +5,33 @@
Author: Piero Tramontano
*)
-section{*The SET Cardholder Registration Protocol*}
+section\<open>The SET Cardholder Registration Protocol\<close>
theory Cardholder_Registration
imports Public_SET
begin
-text{*Note: nonces seem to consist of 20 bytes. That includes both freshness
+text\<open>Note: nonces seem to consist of 20 bytes. That includes both freshness
challenges (Chall-EE, etc.) and important secrets (CardSecret, PANsecret)
-*}
+\<close>
-text{*Simplifications involving @{text analz_image_keys_simps} appear to
+text\<open>Simplifications involving \<open>analz_image_keys_simps\<close> appear to
have become much slower. The cause is unclear. However, there is a big blow-up
-and the rewriting is very sensitive to the set of rewrite rules given.*}
+and the rewriting is very sensitive to the set of rewrite rules given.\<close>
-subsection{*Predicate Formalizing the Encryption Association between Keys *}
+subsection\<open>Predicate Formalizing the Encryption Association between Keys\<close>
primrec KeyCryptKey :: "[key, key, event list] => bool"
where
KeyCryptKey_Nil:
"KeyCryptKey DK K [] = False"
| KeyCryptKey_Cons:
- --{*Says is the only important case.
+ \<comment>\<open>Says is the only important case.
1st case: CR5, where KC3 encrypts KC2.
2nd case: any use of priEK C.
Revision 1.12 has a more complicated version with separate treatment of
the dependency of KC1, KC2 and KC3 on priEK (CA i.) Not needed since
- priEK C is never sent (and so can't be lost except at the start). *}
+ priEK C is never sent (and so can't be lost except at the start).\<close>
"KeyCryptKey DK K (ev # evs) =
(KeyCryptKey DK K evs |
(case ev of
@@ -44,23 +44,23 @@
| Notes A' X => False))"
-subsection{*Predicate formalizing the association between keys and nonces *}
+subsection\<open>Predicate formalizing the association between keys and nonces\<close>
primrec KeyCryptNonce :: "[key, key, event list] => bool"
where
KeyCryptNonce_Nil:
"KeyCryptNonce EK K [] = False"
| KeyCryptNonce_Cons:
- --{*Says is the only important case.
+ \<comment>\<open>Says is the only important case.
1st case: CR3, where KC1 encrypts NC2 (distinct from CR5 due to EXH);
2nd case: CR5, where KC3 encrypts NC3;
3rd case: CR6, where KC2 encrypts NC3;
4th case: CR6, where KC2 encrypts NonceCCA;
5th case: any use of @{term "priEK C"} (including CardSecret).
NB the only Nonces we need to keep secret are CardSecret and NonceCCA.
- But we can't prove @{text Nonce_compromise} unless the relation covers ALL
+ But we can't prove \<open>Nonce_compromise\<close> unless the relation covers ALL
nonces that the protocol keeps secret.
- *}
+\<close>
"KeyCryptNonce DK N (ev # evs) =
(KeyCryptNonce DK N evs |
(case ev of
@@ -85,28 +85,28 @@
| Notes A' X => False))"
-subsection{*Formal protocol definition *}
+subsection\<open>Formal protocol definition\<close>
inductive_set
set_cr :: "event list set"
where
- Nil: --{*Initial trace is empty*}
+ Nil: \<comment>\<open>Initial trace is empty\<close>
"[] \<in> set_cr"
-| Fake: --{*The spy MAY say anything he CAN say.*}
+| Fake: \<comment>\<open>The spy MAY say anything he CAN say.\<close>
"[| evsf \<in> set_cr; X \<in> synth (analz (knows Spy evsf)) |]
==> Says Spy B X # evsf \<in> set_cr"
-| Reception: --{*If A sends a message X to B, then B might receive it*}
+| Reception: \<comment>\<open>If A sends a message X to B, then B might receive it\<close>
"[| evsr \<in> set_cr; Says A B X \<in> set evsr |]
==> Gets B X # evsr \<in> set_cr"
-| SET_CR1: --{*CardCInitReq: C initiates a run, sending a nonce to CCA*}
+| SET_CR1: \<comment>\<open>CardCInitReq: C initiates a run, sending a nonce to CCA\<close>
"[| evs1 \<in> set_cr; C = Cardholder k; Nonce NC1 \<notin> used evs1 |]
==> Says C (CA i) \<lbrace>Agent C, Nonce NC1\<rbrace> # evs1 \<in> set_cr"
-| SET_CR2: --{*CardCInitRes: CA responds sending NC1 and its certificates*}
+| SET_CR2: \<comment>\<open>CardCInitRes: CA responds sending NC1 and its certificates\<close>
"[| evs2 \<in> set_cr;
Gets (CA i) \<lbrace>Agent C, Nonce NC1\<rbrace> \<in> set evs2 |]
==> Says (CA i) C
@@ -116,7 +116,7 @@
# evs2 \<in> set_cr"
| SET_CR3:
- --{*RegFormReq: C sends his PAN and a new nonce to CA.
+ \<comment>\<open>RegFormReq: C sends his PAN and a new nonce to CA.
C verifies that
- nonce received is the same as that sent;
- certificates are signed by RCA;
@@ -126,7 +126,7 @@
checking the signature).
C generates a fresh symmetric key KC1.
The point of encrypting @{term "\<lbrace>Agent C, Nonce NC2, Hash (Pan(pan C))\<rbrace>"}
- is not clear. *}
+ is not clear.\<close>
"[| evs3 \<in> set_cr; C = Cardholder k;
Nonce NC2 \<notin> used evs3;
Key KC1 \<notin> used evs3; KC1 \<in> symKeys;
@@ -140,11 +140,11 @@
# evs3 \<in> set_cr"
| SET_CR4:
- --{*RegFormRes:
+ \<comment>\<open>RegFormRes:
CA responds sending NC2 back with a new nonce NCA, after checking that
- the digital envelope is correctly encrypted by @{term "pubEK (CA i)"}
- the entire message is encrypted with the same key found inside the
- envelope (here, KC1) *}
+ envelope (here, KC1)\<close>
"[| evs4 \<in> set_cr;
Nonce NCA \<notin> used evs4; KC1 \<in> symKeys;
Gets (CA i) (EXHcrypt KC1 EKi \<lbrace>Agent C, Nonce NC2\<rbrace> (Pan(pan X)))
@@ -156,11 +156,11 @@
# evs4 \<in> set_cr"
| SET_CR5:
- --{*CertReq: C sends his PAN, a new nonce, its proposed public signature key
+ \<comment>\<open>CertReq: C sends his PAN, a new nonce, its proposed public signature key
and its half of the secret value to CA.
We now assume that C has a fixed key pair, and he submits (pubSK C).
The protocol does not require this key to be fresh.
- The encryption below is actually EncX.*}
+ The encryption below is actually EncX.\<close>
"[| evs5 \<in> set_cr; C = Cardholder k;
Nonce NC3 \<notin> used evs5; Nonce CardSecret \<notin> used evs5; NC3\<noteq>CardSecret;
Key KC2 \<notin> used evs5; KC2 \<in> symKeys;
@@ -183,13 +183,13 @@
# evs5 \<in> set_cr"
- --{* CertRes: CA responds sending NC3 back with its half of the secret value,
+ \<comment>\<open>CertRes: CA responds sending NC3 back with its half of the secret value,
its signature certificate and the new cardholder signature
certificate. CA checks to have never certified the key proposed by C.
NOTE: In Merchant Registration, the corresponding rule (4)
uses the "sign" primitive. The encryption below is actually @{term EncK},
which is just @{term "Crypt K (sign SK X)"}.
-*}
+\<close>
| SET_CR6:
"[| evs6 \<in> set_cr;
@@ -218,15 +218,15 @@
declare analz_into_parts [dest]
declare Fake_parts_insert_in_Un [dest]
-text{*A "possibility property": there are traces that reach the end.
- An unconstrained proof with many subgoals.*}
+text\<open>A "possibility property": there are traces that reach the end.
+ An unconstrained proof with many subgoals.\<close>
lemma Says_to_Gets:
"Says A B X # evs \<in> set_cr ==> Gets B X # Says A B X # evs \<in> set_cr"
by (rule set_cr.Reception, auto)
-text{*The many nonces and keys generated, some simultaneously, force us to
- introduce them explicitly as shown below.*}
+text\<open>The many nonces and keys generated, some simultaneously, force us to
+ introduce them explicitly as shown below.\<close>
lemma possibility_CR6:
"[|NC1 < (NC2::nat); NC2 < NC3; NC3 < NCA ;
NCA < NonceCCA; NonceCCA < CardSecret;
@@ -262,7 +262,7 @@
apply (simp_all (no_asm_simp) add: symKeys_neq_imp_neq)
done
-text{*General facts about message reception*}
+text\<open>General facts about message reception\<close>
lemma Gets_imp_Says:
"[| Gets B X \<in> set evs; evs \<in> set_cr |] ==> \<exists>A. Says A B X \<in> set evs"
apply (erule rev_mp)
@@ -275,9 +275,9 @@
declare Gets_imp_knows_Spy [THEN parts.Inj, dest]
-subsection{*Proofs on keys *}
+subsection\<open>Proofs on keys\<close>
-text{*Spy never sees an agent's private keys! (unless it's bad at start)*}
+text\<open>Spy never sees an agent's private keys! (unless it's bad at start)\<close>
lemma Spy_see_private_Key [simp]:
"evs \<in> set_cr
@@ -293,8 +293,8 @@
declare Spy_analz_private_Key [THEN [2] rev_iffD1, dest!]
-subsection{*Begin Piero's Theorems on Certificates*}
-text{*Trivial in the current model, where certificates by RCA are secure *}
+subsection\<open>Begin Piero's Theorems on Certificates\<close>
+text\<open>Trivial in the current model, where certificates by RCA are secure\<close>
lemma Crypt_valid_pubEK:
"[| Crypt (priSK RCA) \<lbrace>Agent C, Key EKi, onlyEnc\<rbrace>
@@ -334,7 +334,7 @@
==> EKi = pubEK C & SKi = pubSK C"
by (blast dest: certificate_valid_pubEK certificate_valid_pubSK)
-text{*Nobody can have used non-existent keys!*}
+text\<open>Nobody can have used non-existent keys!\<close>
lemma new_keys_not_used:
"[|K \<in> symKeys; Key K \<notin> used evs; evs \<in> set_cr|]
==> K \<notin> keysFor (parts (knows Spy evs))"
@@ -343,12 +343,12 @@
apply (erule set_cr.induct)
apply (frule_tac [8] Gets_certificate_valid)
apply (frule_tac [6] Gets_certificate_valid, simp_all)
-apply (force dest!: usedI keysFor_parts_insert) --{*Fake*}
-apply (blast,auto) --{*Others*}
+apply (force dest!: usedI keysFor_parts_insert) \<comment>\<open>Fake\<close>
+apply (blast,auto) \<comment>\<open>Others\<close>
done
-subsection{*New versions: as above, but generalized to have the KK argument *}
+subsection\<open>New versions: as above, but generalized to have the KK argument\<close>
lemma gen_new_keys_not_used:
"[|Key K \<notin> used evs; K \<in> symKeys; evs \<in> set_cr |]
@@ -382,10 +382,10 @@
(*<*)
-subsection{*Messages signed by CA*}
+subsection\<open>Messages signed by CA\<close>
-text{*Message @{text SET_CR2}: C can check CA's signature if he has received
- CA's certificate.*}
+text\<open>Message \<open>SET_CR2\<close>: C can check CA's signature if he has received
+ CA's certificate.\<close>
lemma CA_Says_2_lemma:
"[| Crypt (priSK (CA i)) (Hash\<lbrace>Agent C, Nonce NC1\<rbrace>)
\<in> parts (knows Spy evs);
@@ -396,7 +396,7 @@
apply (erule set_cr.induct, auto)
done
-text{*Ever used?*}
+text\<open>Ever used?\<close>
lemma CA_Says_2:
"[| Crypt (invKey SK) (Hash\<lbrace>Agent C, Nonce NC1\<rbrace>)
\<in> parts (knows Spy evs);
@@ -407,8 +407,8 @@
by (blast dest!: certificate_valid_pubSK intro!: CA_Says_2_lemma)
-text{*Message @{text SET_CR4}: C can check CA's signature if he has received
- CA's certificate.*}
+text\<open>Message \<open>SET_CR4\<close>: C can check CA's signature if he has received
+ CA's certificate.\<close>
lemma CA_Says_4_lemma:
"[| Crypt (priSK (CA i)) (Hash\<lbrace>Agent C, Nonce NC2, Nonce NCA\<rbrace>)
\<in> parts (knows Spy evs);
@@ -419,7 +419,7 @@
apply (erule set_cr.induct, auto)
done
-text{*NEVER USED*}
+text\<open>NEVER USED\<close>
lemma CA_Says_4:
"[| Crypt (invKey SK) (Hash\<lbrace>Agent C, Nonce NC2, Nonce NCA\<rbrace>)
\<in> parts (knows Spy evs);
@@ -430,8 +430,8 @@
by (blast dest!: certificate_valid_pubSK intro!: CA_Says_4_lemma)
-text{*Message @{text SET_CR6}: C can check CA's signature if he has
- received CA's certificate.*}
+text\<open>Message \<open>SET_CR6\<close>: C can check CA's signature if he has
+ received CA's certificate.\<close>
lemma CA_Says_6_lemma:
"[| Crypt (priSK (CA i))
(Hash\<lbrace>Agent C, Nonce NC3, Agent (CA i), Nonce NonceCCA\<rbrace>)
@@ -443,7 +443,7 @@
apply (erule set_cr.induct, auto)
done
-text{*NEVER USED*}
+text\<open>NEVER USED\<close>
lemma CA_Says_6:
"[| Crypt (invKey SK) (Hash\<lbrace>Agent C, Nonce NC3, Agent (CA i), Nonce NonceCCA\<rbrace>)
\<in> parts (knows Spy evs);
@@ -455,17 +455,17 @@
(*>*)
-subsection{*Useful lemmas *}
+subsection\<open>Useful lemmas\<close>
-text{*Rewriting rule for private encryption keys. Analogous rewriting rules
-for other keys aren't needed.*}
+text\<open>Rewriting rule for private encryption keys. Analogous rewriting rules
+for other keys aren't needed.\<close>
lemma parts_image_priEK:
"[|Key (priEK C) \<in> parts (Key`KK Un (knows Spy evs));
evs \<in> set_cr|] ==> priEK C \<in> KK | C \<in> bad"
by auto
-text{*trivial proof because (priEK C) never appears even in (parts evs)*}
+text\<open>trivial proof because (priEK C) never appears even in (parts evs)\<close>
lemma analz_image_priEK:
"evs \<in> set_cr ==>
(Key (priEK C) \<in> analz (Key`KK Un (knows Spy evs))) =
@@ -473,12 +473,12 @@
by (blast dest!: parts_image_priEK intro: analz_mono [THEN [2] rev_subsetD])
-subsection{*Secrecy of Session Keys *}
+subsection\<open>Secrecy of Session Keys\<close>
-subsubsection{*Lemmas about the predicate KeyCryptKey *}
+subsubsection\<open>Lemmas about the predicate KeyCryptKey\<close>
-text{*A fresh DK cannot be associated with any other
- (with respect to a given trace). *}
+text\<open>A fresh DK cannot be associated with any other
+ (with respect to a given trace).\<close>
lemma DK_fresh_not_KeyCryptKey:
"[| Key DK \<notin> used evs; evs \<in> set_cr |] ==> ~ KeyCryptKey DK K evs"
apply (erule rev_mp)
@@ -487,9 +487,9 @@
apply (blast dest: Crypt_analz_imp_used)+
done
-text{*A fresh K cannot be associated with any other. The assumption that
+text\<open>A fresh K cannot be associated with any other. The assumption that
DK isn't a private encryption key may be an artifact of the particular
- definition of KeyCryptKey.*}
+ definition of KeyCryptKey.\<close>
lemma K_fresh_not_KeyCryptKey:
"[|\<forall>C. DK \<noteq> priEK C; Key K \<notin> used evs|] ==> ~ KeyCryptKey DK K evs"
apply (induct evs)
@@ -497,8 +497,8 @@
done
-text{*This holds because if (priEK (CA i)) appears in any traffic then it must
- be known to the Spy, by @{term Spy_see_private_Key}*}
+text\<open>This holds because if (priEK (CA i)) appears in any traffic then it must
+ be known to the Spy, by @{term Spy_see_private_Key}\<close>
lemma cardSK_neq_priEK:
"[|Key cardSK \<notin> analz (knows Spy evs);
Key cardSK : parts (knows Spy evs);
@@ -510,16 +510,16 @@
Key cardSK \<notin> analz (knows Spy evs) --> ~ KeyCryptKey cardSK K evs"
by (erule set_cr.induct, analz_mono_contra, auto)
-text{*Lemma for message 5: pubSK C is never used to encrypt Keys.*}
+text\<open>Lemma for message 5: pubSK C is never used to encrypt Keys.\<close>
lemma pubSK_not_KeyCryptKey [simp]: "~ KeyCryptKey (pubSK C) K evs"
apply (induct_tac "evs")
apply (auto simp add: parts_insert2 split add: event.split)
done
-text{*Lemma for message 6: either cardSK is compromised (when we don't care)
+text\<open>Lemma for message 6: either cardSK is compromised (when we don't care)
or else cardSK hasn't been used to encrypt K. Previously we treated
message 5 in the same way, but the current model assumes that rule
- @{text SET_CR5} is executed only by honest agents.*}
+ \<open>SET_CR5\<close> is executed only by honest agents.\<close>
lemma msg6_KeyCryptKey_disj:
"[|Gets B \<lbrace>Crypt KC3 \<lbrace>Agent C, Nonce N, Key KC2, Key cardSK, X\<rbrace>, Y\<rbrace>
\<in> set evs;
@@ -528,23 +528,23 @@
(\<forall>K. ~ KeyCryptKey cardSK K evs)"
by (blast dest: not_KeyCryptKey_cardSK intro: cardSK_neq_priEK)
-text{*As usual: we express the property as a logical equivalence*}
+text\<open>As usual: we express the property as a logical equivalence\<close>
lemma Key_analz_image_Key_lemma:
"P --> (Key K \<in> analz (Key`KK Un H)) --> (K \<in> KK | Key K \<in> analz H)
==>
P --> (Key K \<in> analz (Key`KK Un H)) = (K \<in> KK | Key K \<in> analz H)"
by (blast intro: analz_mono [THEN [2] rev_subsetD])
-method_setup valid_certificate_tac = {*
+method_setup valid_certificate_tac = \<open>
Args.goal_spec >> (fn quant => fn ctxt => SIMPLE_METHOD'' quant
(fn i =>
EVERY [forward_tac ctxt @{thms Gets_certificate_valid} i,
assume_tac ctxt i,
eresolve_tac ctxt [conjE] i, REPEAT (hyp_subst_tac ctxt i)]))
-*}
+\<close>
-text{*The @{text "(no_asm)"} attribute is essential, since it retains
- the quantifier and allows the simprule's condition to itself be simplified.*}
+text\<open>The \<open>(no_asm)\<close> attribute is essential, since it retains
+ the quantifier and allows the simprule's condition to itself be simplified.\<close>
lemma symKey_compromise [rule_format (no_asm)]:
"evs \<in> set_cr ==>
(\<forall>SK KK. SK \<in> symKeys \<longrightarrow> (\<forall>K \<in> KK. ~ KeyCryptKey K SK evs) -->
@@ -553,8 +553,8 @@
apply (erule set_cr.induct)
apply (rule_tac [!] allI) +
apply (rule_tac [!] impI [THEN Key_analz_image_Key_lemma, THEN impI])+
-apply (valid_certificate_tac [8]) --{*for message 5*}
-apply (valid_certificate_tac [6]) --{*for message 5*}
+apply (valid_certificate_tac [8]) \<comment>\<open>for message 5\<close>
+apply (valid_certificate_tac [6]) \<comment>\<open>for message 5\<close>
apply (erule_tac [9] msg6_KeyCryptKey_disj [THEN disjE])
apply (simp_all
del: image_insert image_Un imp_disjL
@@ -563,15 +563,15 @@
K_fresh_not_KeyCryptKey
DK_fresh_not_KeyCryptKey ball_conj_distrib
analz_image_priEK disj_simps)
- --{*9 seconds on a 1.6GHz machine*}
+ \<comment>\<open>9 seconds on a 1.6GHz machine\<close>
apply spy_analz
-apply blast --{*3*}
-apply blast --{*5*}
+apply blast \<comment>\<open>3\<close>
+apply blast \<comment>\<open>5\<close>
done
-text{*The remaining quantifiers seem to be essential.
+text\<open>The remaining quantifiers seem to be essential.
NO NEED to assume the cardholder's OK: bad cardholders don't do anything
- wrong!!*}
+ wrong!!\<close>
lemma symKey_secrecy [rule_format]:
"[|CA i \<notin> bad; K \<in> symKeys; evs \<in> set_cr|]
==> \<forall>X c. Says (Cardholder c) (CA i) X \<in> set evs -->
@@ -579,8 +579,8 @@
Cardholder c \<notin> bad -->
Key K \<notin> analz (knows Spy evs)"
apply (erule set_cr.induct)
-apply (frule_tac [8] Gets_certificate_valid) --{*for message 5*}
-apply (frule_tac [6] Gets_certificate_valid) --{*for message 3*}
+apply (frule_tac [8] Gets_certificate_valid) \<comment>\<open>for message 5\<close>
+apply (frule_tac [6] Gets_certificate_valid) \<comment>\<open>for message 3\<close>
apply (erule_tac [11] msg6_KeyCryptKey_disj [THEN disjE])
apply (simp_all del: image_insert image_Un imp_disjL
add: symKey_compromise fresh_notin_analz_knows_Spy
@@ -589,18 +589,18 @@
K_fresh_not_KeyCryptKey
DK_fresh_not_KeyCryptKey
analz_image_priEK)
- --{*2.5 seconds on a 1.6GHz machine*}
-apply spy_analz --{*Fake*}
+ \<comment>\<open>2.5 seconds on a 1.6GHz machine\<close>
+apply spy_analz \<comment>\<open>Fake\<close>
apply (auto intro: analz_into_parts [THEN usedI] in_parts_Says_imp_used)
done
-subsection{*Primary Goals of Cardholder Registration *}
+subsection\<open>Primary Goals of Cardholder Registration\<close>
-text{*The cardholder's certificate really was created by the CA, provided the
- CA is uncompromised *}
+text\<open>The cardholder's certificate really was created by the CA, provided the
+ CA is uncompromised\<close>
-text{*Lemma concerning the actual signed message digest*}
+text\<open>Lemma concerning the actual signed message digest\<close>
lemma cert_valid_lemma:
"[|Crypt (priSK (CA i)) \<lbrace>Hash \<lbrace>Nonce N, Pan(pan C)\<rbrace>, Key cardSK, N1\<rbrace>
\<in> parts (knows Spy evs);
@@ -615,8 +615,8 @@
apply auto
done
-text{*Pre-packaged version for cardholder. We don't try to confirm the values
- of KC2, X and Y, since they are not important.*}
+text\<open>Pre-packaged version for cardholder. We don't try to confirm the values
+ of KC2, X and Y, since they are not important.\<close>
lemma certificate_valid_cardSK:
"[|Gets C (Crypt KC2 \<lbrace>X, certC (pan C) cardSK N onlySig (invKey SKi),
cert (CA i) SKi onlySig (priSK RCA)\<rbrace>) \<in> set evs;
@@ -648,12 +648,12 @@
done
-subsection{*Secrecy of Nonces*}
+subsection\<open>Secrecy of Nonces\<close>
-subsubsection{*Lemmas about the predicate KeyCryptNonce *}
+subsubsection\<open>Lemmas about the predicate KeyCryptNonce\<close>
-text{*A fresh DK cannot be associated with any other
- (with respect to a given trace). *}
+text\<open>A fresh DK cannot be associated with any other
+ (with respect to a given trace).\<close>
lemma DK_fresh_not_KeyCryptNonce:
"[| DK \<in> symKeys; Key DK \<notin> used evs; evs \<in> set_cr |]
==> ~ KeyCryptNonce DK K evs"
@@ -666,8 +666,8 @@
apply (auto simp add: DK_fresh_not_KeyCryptKey)
done
-text{*A fresh N cannot be associated with any other
- (with respect to a given trace). *}
+text\<open>A fresh N cannot be associated with any other
+ (with respect to a given trace).\<close>
lemma N_fresh_not_KeyCryptNonce:
"\<forall>C. DK \<noteq> priEK C ==> Nonce N \<notin> used evs --> ~ KeyCryptNonce DK N evs"
apply (induct_tac "evs")
@@ -680,21 +680,21 @@
"[|cardSK \<notin> symKeys; \<forall>C. cardSK \<noteq> priEK C; evs \<in> set_cr|] ==>
Key cardSK \<notin> analz (knows Spy evs) --> ~ KeyCryptNonce cardSK N evs"
apply (erule set_cr.induct, analz_mono_contra, simp_all)
-apply (blast dest: not_KeyCryptKey_cardSK) --{*6*}
+apply (blast dest: not_KeyCryptKey_cardSK) \<comment>\<open>6\<close>
done
-subsubsection{*Lemmas for message 5 and 6:
+subsubsection\<open>Lemmas for message 5 and 6:
either cardSK is compromised (when we don't care)
- or else cardSK hasn't been used to encrypt K. *}
+ or else cardSK hasn't been used to encrypt K.\<close>
-text{*Lemma for message 5: pubSK C is never used to encrypt Nonces.*}
+text\<open>Lemma for message 5: pubSK C is never used to encrypt Nonces.\<close>
lemma pubSK_not_KeyCryptNonce [simp]: "~ KeyCryptNonce (pubSK C) N evs"
apply (induct_tac "evs")
apply (auto simp add: parts_insert2 split add: event.split)
done
-text{*Lemma for message 6: either cardSK is compromised (when we don't care)
- or else cardSK hasn't been used to encrypt K.*}
+text\<open>Lemma for message 6: either cardSK is compromised (when we don't care)
+ or else cardSK hasn't been used to encrypt K.\<close>
lemma msg6_KeyCryptNonce_disj:
"[|Gets B \<lbrace>Crypt KC3 \<lbrace>Agent C, Nonce N, Key KC2, Key cardSK, X\<rbrace>, Y\<rbrace>
\<in> set evs;
@@ -706,15 +706,15 @@
intro: cardSK_neq_priEK)
-text{*As usual: we express the property as a logical equivalence*}
+text\<open>As usual: we express the property as a logical equivalence\<close>
lemma Nonce_analz_image_Key_lemma:
"P --> (Nonce N \<in> analz (Key`KK Un H)) --> (Nonce N \<in> analz H)
==> P --> (Nonce N \<in> analz (Key`KK Un H)) = (Nonce N \<in> analz H)"
by (blast intro: analz_mono [THEN [2] rev_subsetD])
-text{*The @{text "(no_asm)"} attribute is essential, since it retains
- the quantifier and allows the simprule's condition to itself be simplified.*}
+text\<open>The \<open>(no_asm)\<close> attribute is essential, since it retains
+ the quantifier and allows the simprule's condition to itself be simplified.\<close>
lemma Nonce_compromise [rule_format (no_asm)]:
"evs \<in> set_cr ==>
(\<forall>N KK. (\<forall>K \<in> KK. ~ KeyCryptNonce K N evs) -->
@@ -723,8 +723,8 @@
apply (erule set_cr.induct)
apply (rule_tac [!] allI)+
apply (rule_tac [!] impI [THEN Nonce_analz_image_Key_lemma])+
-apply (frule_tac [8] Gets_certificate_valid) --{*for message 5*}
-apply (frule_tac [6] Gets_certificate_valid) --{*for message 3*}
+apply (frule_tac [8] Gets_certificate_valid) \<comment>\<open>for message 5\<close>
+apply (frule_tac [6] Gets_certificate_valid) \<comment>\<open>for message 3\<close>
apply (frule_tac [11] msg6_KeyCryptNonce_disj)
apply (erule_tac [13] disjE)
apply (simp_all del: image_insert image_Un
@@ -734,21 +734,21 @@
N_fresh_not_KeyCryptNonce
DK_fresh_not_KeyCryptNonce K_fresh_not_KeyCryptKey
ball_conj_distrib analz_image_priEK)
- --{*14 seconds on a 1.6GHz machine*}
-apply spy_analz --{*Fake*}
-apply blast --{*3*}
-apply blast --{*5*}
-txt{*Message 6*}
+ \<comment>\<open>14 seconds on a 1.6GHz machine\<close>
+apply spy_analz \<comment>\<open>Fake\<close>
+apply blast \<comment>\<open>3\<close>
+apply blast \<comment>\<open>5\<close>
+txt\<open>Message 6\<close>
apply (metis symKey_compromise)
- --{*cardSK compromised*}
-txt{*Simplify again--necessary because the previous simplification introduces
- some logical connectives*}
+ \<comment>\<open>cardSK compromised\<close>
+txt\<open>Simplify again--necessary because the previous simplification introduces
+ some logical connectives\<close>
apply (force simp del: image_insert image_Un imp_disjL
simp add: analz_image_keys_simps symKey_compromise)
done
-subsection{*Secrecy of CardSecret: the Cardholder's secret*}
+subsection\<open>Secrecy of CardSecret: the Cardholder's secret\<close>
lemma NC2_not_CardSecret:
"[|Crypt EKj \<lbrace>Key K, Pan p, Hash \<lbrace>Agent D, Nonce N\<rbrace>\<rbrace>
@@ -773,12 +773,12 @@
Cardholder k \<notin> bad & CA i \<notin> bad)"
apply (erule_tac P = "U \<in> H" for H in rev_mp)
apply (erule set_cr.induct)
-apply (valid_certificate_tac [8]) --{*for message 5*}
+apply (valid_certificate_tac [8]) \<comment>\<open>for message 5\<close>
apply (simp_all del: image_insert image_Un imp_disjL
add: analz_image_keys_simps analz_knows_absorb
analz_knows_absorb2 notin_image_iff)
- --{*4 seconds on a 1.6GHz machine*}
-apply (simp_all (no_asm_simp)) --{*leaves 4 subgoals*}
+ \<comment>\<open>4 seconds on a 1.6GHz machine\<close>
+apply (simp_all (no_asm_simp)) \<comment>\<open>leaves 4 subgoals\<close>
apply (blast intro!: analz_insertI)+
done
@@ -790,7 +790,7 @@
by (force dest!: refl [THEN KC2_secure_lemma] symKey_secrecy)
-text{*Inductive version*}
+text\<open>Inductive version\<close>
lemma CardSecret_secrecy_lemma [rule_format]:
"[|CA i \<notin> bad; evs \<in> set_cr|]
==> Key K \<notin> analz (knows Spy evs) -->
@@ -798,8 +798,8 @@
\<in> parts (knows Spy evs) -->
Nonce CardSecret \<notin> analz (knows Spy evs)"
apply (erule set_cr.induct, analz_mono_contra)
-apply (valid_certificate_tac [8]) --{*for message 5*}
-apply (valid_certificate_tac [6]) --{*for message 5*}
+apply (valid_certificate_tac [8]) \<comment>\<open>for message 5\<close>
+apply (valid_certificate_tac [6]) \<comment>\<open>for message 5\<close>
apply (frule_tac [9] msg6_KeyCryptNonce_disj [THEN disjE])
apply (simp_all
del: image_insert image_Un imp_disjL
@@ -809,19 +809,19 @@
N_fresh_not_KeyCryptNonce DK_fresh_not_KeyCryptNonce
ball_conj_distrib Nonce_compromise symKey_compromise
analz_image_priEK)
- --{*2.5 seconds on a 1.6GHz machine*}
-apply spy_analz --{*Fake*}
+ \<comment>\<open>2.5 seconds on a 1.6GHz machine\<close>
+apply spy_analz \<comment>\<open>Fake\<close>
apply (simp_all (no_asm_simp))
-apply blast --{*1*}
-apply (blast dest!: Gets_imp_knows_Spy [THEN analz.Inj]) --{*2*}
-apply blast --{*3*}
-apply (blast dest: NC2_not_CardSecret Gets_imp_knows_Spy [THEN analz.Inj] analz_symKeys_Decrypt) --{*4*}
-apply blast --{*5*}
-apply (blast dest: KC2_secrecy)+ --{*Message 6: two cases*}
+apply blast \<comment>\<open>1\<close>
+apply (blast dest!: Gets_imp_knows_Spy [THEN analz.Inj]) \<comment>\<open>2\<close>
+apply blast \<comment>\<open>3\<close>
+apply (blast dest: NC2_not_CardSecret Gets_imp_knows_Spy [THEN analz.Inj] analz_symKeys_Decrypt) \<comment>\<open>4\<close>
+apply blast \<comment>\<open>5\<close>
+apply (blast dest: KC2_secrecy)+ \<comment>\<open>Message 6: two cases\<close>
done
-text{*Packaged version for cardholder*}
+text\<open>Packaged version for cardholder\<close>
lemma CardSecret_secrecy:
"[|Cardholder k \<notin> bad; CA i \<notin> bad;
Says (Cardholder k) (CA i)
@@ -838,7 +838,7 @@
done
-subsection{*Secrecy of NonceCCA [the CA's secret] *}
+subsection\<open>Secrecy of NonceCCA [the CA's secret]\<close>
lemma NC2_not_NonceCCA:
"[|Hash \<lbrace>Agent C', Nonce N', Agent C, Nonce N\<rbrace>
@@ -853,7 +853,7 @@
done
-text{*Inductive version*}
+text\<open>Inductive version\<close>
lemma NonceCCA_secrecy_lemma [rule_format]:
"[|CA i \<notin> bad; evs \<in> set_cr|]
==> Key K \<notin> analz (knows Spy evs) -->
@@ -864,8 +864,8 @@
\<in> parts (knows Spy evs) -->
Nonce NonceCCA \<notin> analz (knows Spy evs)"
apply (erule set_cr.induct, analz_mono_contra)
-apply (valid_certificate_tac [8]) --{*for message 5*}
-apply (valid_certificate_tac [6]) --{*for message 5*}
+apply (valid_certificate_tac [8]) \<comment>\<open>for message 5\<close>
+apply (valid_certificate_tac [6]) \<comment>\<open>for message 5\<close>
apply (frule_tac [9] msg6_KeyCryptNonce_disj [THEN disjE])
apply (simp_all
del: image_insert image_Un imp_disjL
@@ -875,18 +875,18 @@
N_fresh_not_KeyCryptNonce DK_fresh_not_KeyCryptNonce
ball_conj_distrib Nonce_compromise symKey_compromise
analz_image_priEK)
- --{*3 seconds on a 1.6GHz machine*}
-apply spy_analz --{*Fake*}
-apply blast --{*1*}
-apply (blast dest!: Gets_imp_knows_Spy [THEN analz.Inj]) --{*2*}
-apply blast --{*3*}
-apply (blast dest: NC2_not_NonceCCA) --{*4*}
-apply blast --{*5*}
-apply (blast dest: KC2_secrecy)+ --{*Message 6: two cases*}
+ \<comment>\<open>3 seconds on a 1.6GHz machine\<close>
+apply spy_analz \<comment>\<open>Fake\<close>
+apply blast \<comment>\<open>1\<close>
+apply (blast dest!: Gets_imp_knows_Spy [THEN analz.Inj]) \<comment>\<open>2\<close>
+apply blast \<comment>\<open>3\<close>
+apply (blast dest: NC2_not_NonceCCA) \<comment>\<open>4\<close>
+apply blast \<comment>\<open>5\<close>
+apply (blast dest: KC2_secrecy)+ \<comment>\<open>Message 6: two cases\<close>
done
-text{*Packaged version for cardholder*}
+text\<open>Packaged version for cardholder\<close>
lemma NonceCCA_secrecy:
"[|Cardholder k \<notin> bad; CA i \<notin> bad;
Gets (Cardholder k)
@@ -906,16 +906,16 @@
apply (auto simp add: parts_insert2)
done
-text{*We don't bother to prove guarantees for the CA. He doesn't care about
- the PANSecret: it isn't his credit card!*}
+text\<open>We don't bother to prove guarantees for the CA. He doesn't care about
+ the PANSecret: it isn't his credit card!\<close>
-subsection{*Rewriting Rule for PANs*}
+subsection\<open>Rewriting Rule for PANs\<close>
-text{*Lemma for message 6: either cardSK isn't a CA's private encryption key,
+text\<open>Lemma for message 6: either cardSK isn't a CA's private encryption key,
or if it is then (because it appears in traffic) that CA is bad,
and so the Spy knows that key already. Either way, we can simplify
- the expression @{term "analz (insert (Key cardSK) X)"}.*}
+ the expression @{term "analz (insert (Key cardSK) X)"}.\<close>
lemma msg6_cardSK_disj:
"[|Gets A \<lbrace>Crypt K \<lbrace>c, n, k', Key cardSK, X\<rbrace>, Y\<rbrace>
\<in> set evs; evs \<in> set_cr |]
@@ -935,16 +935,16 @@
apply (erule set_cr.induct)
apply (rule_tac [!] allI impI)+
apply (rule_tac [!] analz_image_pan_lemma)
-apply (valid_certificate_tac [8]) --{*for message 5*}
-apply (valid_certificate_tac [6]) --{*for message 5*}
+apply (valid_certificate_tac [8]) \<comment>\<open>for message 5\<close>
+apply (valid_certificate_tac [6]) \<comment>\<open>for message 5\<close>
apply (erule_tac [9] msg6_cardSK_disj [THEN disjE])
apply (simp_all
del: image_insert image_Un
add: analz_image_keys_simps disjoint_image_iff
notin_image_iff analz_image_priEK)
- --{*6 seconds on a 1.6GHz machine*}
+ \<comment>\<open>6 seconds on a 1.6GHz machine\<close>
apply spy_analz
-apply (simp add: insert_absorb) --{*6*}
+apply (simp add: insert_absorb) \<comment>\<open>6\<close>
done
lemma analz_insert_pan:
@@ -955,9 +955,9 @@
add: analz_image_keys_simps analz_image_pan)
-text{*Confidentiality of the PAN\@. Maybe we could combine the statements of
+text\<open>Confidentiality of the PAN\@. Maybe we could combine the statements of
this theorem with @{term analz_image_pan}, requiring a single induction but
- a much more difficult proof.*}
+ a much more difficult proof.\<close>
lemma pan_confidentiality:
"[| Pan (pan C) \<in> analz(knows Spy evs); C \<noteq>Spy; evs :set_cr|]
==> \<exists>i X K HN.
@@ -966,22 +966,22 @@
& (CA i) \<in> bad"
apply (erule rev_mp)
apply (erule set_cr.induct)
-apply (valid_certificate_tac [8]) --{*for message 5*}
-apply (valid_certificate_tac [6]) --{*for message 5*}
+apply (valid_certificate_tac [8]) \<comment>\<open>for message 5\<close>
+apply (valid_certificate_tac [6]) \<comment>\<open>for message 5\<close>
apply (erule_tac [9] msg6_cardSK_disj [THEN disjE])
apply (simp_all
del: image_insert image_Un
add: analz_image_keys_simps analz_insert_pan analz_image_pan
notin_image_iff analz_image_priEK)
- --{*3.5 seconds on a 1.6GHz machine*}
-apply spy_analz --{*fake*}
-apply blast --{*3*}
-apply blast --{*5*}
-apply (simp (no_asm_simp) add: insert_absorb) --{*6*}
+ \<comment>\<open>3.5 seconds on a 1.6GHz machine\<close>
+apply spy_analz \<comment>\<open>fake\<close>
+apply blast \<comment>\<open>3\<close>
+apply blast \<comment>\<open>5\<close>
+apply (simp (no_asm_simp) add: insert_absorb) \<comment>\<open>6\<close>
done
-subsection{*Unicity*}
+subsection\<open>Unicity\<close>
lemma CR6_Says_imp_Notes:
"[|Says (CA i) C (Crypt KC2
@@ -995,8 +995,8 @@
apply (simp_all (no_asm_simp))
done
-text{*Unicity of cardSK: it uniquely identifies the other components.
- This holds because a CA accepts a cardSK at most once.*}
+text\<open>Unicity of cardSK: it uniquely identifies the other components.
+ This holds because a CA accepts a cardSK at most once.\<close>
lemma cardholder_key_unicity:
"[|Says (CA i) C (Crypt KC2
\<lbrace>sign (priSK (CA i)) \<lbrace>Agent C, Nonce NC3, Agent (CA i), Nonce Y\<rbrace>,
@@ -1018,7 +1018,7 @@
(*<*)
-text{*UNUSED unicity result*}
+text\<open>UNUSED unicity result\<close>
lemma unique_KC1:
"[|Says C B \<lbrace>Crypt KC1 X, Crypt EK \<lbrace>Key KC1, Y\<rbrace>\<rbrace>
\<in> set evs;
@@ -1030,7 +1030,7 @@
apply (erule set_cr.induct, auto)
done
-text{*UNUSED unicity result*}
+text\<open>UNUSED unicity result\<close>
lemma unique_KC2:
"[|Says C B \<lbrace>Crypt K \<lbrace>Agent C, nn, Key KC2, X\<rbrace>, Y\<rbrace> \<in> set evs;
Says C B' \<lbrace>Crypt K' \<lbrace>Agent C, nn', Key KC2, X'\<rbrace>, Y'\<rbrace> \<in> set evs;
@@ -1042,8 +1042,8 @@
(*>*)
-text{*Cannot show cardSK to be secret because it isn't assumed to be fresh
- it could be a previously compromised cardSK [e.g. involving a bad CA]*}
+text\<open>Cannot show cardSK to be secret because it isn't assumed to be fresh
+ it could be a previously compromised cardSK [e.g. involving a bad CA]\<close>
end
--- a/src/HOL/SET_Protocol/Event_SET.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/SET_Protocol/Event_SET.thy Thu May 26 17:51:22 2016 +0200
@@ -4,34 +4,34 @@
Author: Lawrence C Paulson
*)
-section{*Theory of Events for SET*}
+section\<open>Theory of Events for SET\<close>
theory Event_SET
imports Message_SET
begin
-text{*The Root Certification Authority*}
+text\<open>The Root Certification Authority\<close>
abbreviation "RCA == CA 0"
-text{*Message events*}
+text\<open>Message events\<close>
datatype
event = Says agent agent msg
| Gets agent msg
| Notes agent msg
-text{*compromised agents: keys known, Notes visible*}
+text\<open>compromised agents: keys known, Notes visible\<close>
consts bad :: "agent set"
-text{*Spy has access to his own key for spoof messages, but RCA is secure*}
+text\<open>Spy has access to his own key for spoof messages, but RCA is secure\<close>
specification (bad)
Spy_in_bad [iff]: "Spy \<in> bad"
RCA_not_bad [iff]: "RCA \<notin> bad"
by (rule exI [of _ "{Spy}"], simp)
-subsection{*Agents' Knowledge*}
+subsection\<open>Agents' Knowledge\<close>
consts (*Initial states of agents -- parameter of the construction*)
initState :: "agent => msg set"
@@ -60,7 +60,7 @@
if A'=A then insert X (knows A evs) else knows A evs))"
-subsection{*Used Messages*}
+subsection\<open>Used Messages\<close>
primrec used :: "event list => msg set"
where
@@ -88,8 +88,8 @@
"knows Spy (Says A B X # evs) = insert X (knows Spy evs)"
by auto
-text{*Letting the Spy see "bad" agents' notes avoids redundant case-splits
- on whether @{term "A=Spy"} and whether @{term "A\<in>bad"}*}
+text\<open>Letting the Spy see "bad" agents' notes avoids redundant case-splits
+ on whether @{term "A=Spy"} and whether @{term "A\<in>bad"}\<close>
lemma knows_Spy_Notes [simp]:
"knows Spy (Notes A X # evs) =
(if A:bad then insert X (knows Spy evs) else knows Spy evs)"
@@ -130,7 +130,7 @@
parts.Body [elim_format]
-subsection{*The Function @{term used}*}
+subsection\<open>The Function @{term used}\<close>
lemma parts_knows_Spy_subset_used: "parts (knows Spy evs) <= used evs"
apply (induct_tac "evs")
@@ -163,15 +163,15 @@
apply (induct_tac [2] "a", auto)
done
-text{*NOTE REMOVAL--laws above are cleaner, as they don't involve "case"*}
+text\<open>NOTE REMOVAL--laws above are cleaner, as they don't involve "case"\<close>
declare knows_Cons [simp del]
used_Nil [simp del] used_Cons [simp del]
-text{*For proving theorems of the form @{term "X \<notin> analz (knows Spy evs) --> P"}
+text\<open>For proving theorems of the form @{term "X \<notin> analz (knows Spy evs) --> P"}
New events added by induction to "evs" are discarded. Provided
this information isn't needed, the proof will be much shorter, since
- it will omit complicated reasoning about @{term analz}.*}
+ it will omit complicated reasoning about @{term analz}.\<close>
lemmas analz_mono_contra =
knows_Spy_subset_knows_Spy_Says [THEN analz_mono, THEN contra_subsetD]
@@ -181,15 +181,15 @@
lemmas analz_impI = impI [where P = "Y \<notin> analz (knows Spy evs)"] for Y evs
ML
-{*
+\<open>
fun analz_mono_contra_tac ctxt =
resolve_tac ctxt @{thms analz_impI} THEN'
REPEAT1 o (dresolve_tac ctxt @{thms analz_mono_contra})
THEN' mp_tac ctxt
-*}
+\<close>
-method_setup analz_mono_contra = {*
- Scan.succeed (fn ctxt => SIMPLE_METHOD (REPEAT_FIRST (analz_mono_contra_tac ctxt))) *}
+method_setup analz_mono_contra = \<open>
+ Scan.succeed (fn ctxt => SIMPLE_METHOD (REPEAT_FIRST (analz_mono_contra_tac ctxt)))\<close>
"for proving theorems of the form X \<notin> analz (knows Spy evs) --> P"
end
--- a/src/HOL/SET_Protocol/Merchant_Registration.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/SET_Protocol/Merchant_Registration.thy Thu May 26 17:51:22 2016 +0200
@@ -4,43 +4,43 @@
Author: Lawrence C Paulson
*)
-section{*The SET Merchant Registration Protocol*}
+section\<open>The SET Merchant Registration Protocol\<close>
theory Merchant_Registration
imports Public_SET
begin
-text{*Copmpared with Cardholder Reigstration, @{text KeyCryptKey} is not
+text\<open>Copmpared with Cardholder Reigstration, \<open>KeyCryptKey\<close> is not
needed: no session key encrypts another. Instead we
prove the "key compromise" theorems for sets KK that contain no private
- encryption keys (@{term "priEK C"}). *}
+ encryption keys (@{term "priEK C"}).\<close>
inductive_set
set_mr :: "event list set"
where
- Nil: --{*Initial trace is empty*}
+ Nil: \<comment>\<open>Initial trace is empty\<close>
"[] \<in> set_mr"
-| Fake: --{*The spy MAY say anything he CAN say.*}
+| Fake: \<comment>\<open>The spy MAY say anything he CAN say.\<close>
"[| evsf \<in> set_mr; X \<in> synth (analz (knows Spy evsf)) |]
==> Says Spy B X # evsf \<in> set_mr"
-| Reception: --{*If A sends a message X to B, then B might receive it*}
+| Reception: \<comment>\<open>If A sends a message X to B, then B might receive it\<close>
"[| evsr \<in> set_mr; Says A B X \<in> set evsr |]
==> Gets B X # evsr \<in> set_mr"
-| SET_MR1: --{*RegFormReq: M requires a registration form to a CA*}
+| SET_MR1: \<comment>\<open>RegFormReq: M requires a registration form to a CA\<close>
"[| evs1 \<in> set_mr; M = Merchant k; Nonce NM1 \<notin> used evs1 |]
==> Says M (CA i) \<lbrace>Agent M, Nonce NM1\<rbrace> # evs1 \<in> set_mr"
-| SET_MR2: --{*RegFormRes: CA replies with the registration form and the
- certificates for her keys*}
+| SET_MR2: \<comment>\<open>RegFormRes: CA replies with the registration form and the
+ certificates for her keys\<close>
"[| evs2 \<in> set_mr; Nonce NCA \<notin> used evs2;
Gets (CA i) \<lbrace>Agent M, Nonce NM1\<rbrace> \<in> set evs2 |]
==> Says (CA i) M \<lbrace>sign (priSK (CA i)) \<lbrace>Agent M, Nonce NM1, Nonce NCA\<rbrace>,
@@ -49,12 +49,12 @@
# evs2 \<in> set_mr"
| SET_MR3:
- --{*CertReq: M submits the key pair to be certified. The Notes
+ \<comment>\<open>CertReq: M submits the key pair to be certified. The Notes
event allows KM1 to be lost if M is compromised. Piero remarks
that the agent mentioned inside the signature is not verified to
correspond to M. As in CR, each Merchant has fixed key pairs. M
is only optionally required to send NCA back, so M doesn't do so
- in the model*}
+ in the model\<close>
"[| evs3 \<in> set_mr; M = Merchant k; Nonce NM2 \<notin> used evs3;
Key KM1 \<notin> used evs3; KM1 \<in> symKeys;
Gets M \<lbrace>sign (invKey SKi) \<lbrace>Agent X, Nonce NM1, Nonce NCA\<rbrace>,
@@ -70,12 +70,12 @@
# evs3 \<in> set_mr"
| SET_MR4:
- --{*CertRes: CA issues the certificates for merSK and merEK,
+ \<comment>\<open>CertRes: CA issues the certificates for merSK and merEK,
while checking never to have certified the m even
separately. NOTE: In Cardholder Registration the
corresponding rule (6) doesn't use the "sign" primitive. "The
CertRes shall be signed but not encrypted if the EE is a Merchant
- or Payment Gateway."-- Programmer's Guide, page 191.*}
+ or Payment Gateway."-- Programmer's Guide, page 191.\<close>
"[| evs4 \<in> set_mr; M = Merchant k;
merSK \<notin> symKeys; merEK \<notin> symKeys;
Notes (CA i) (Key merSK) \<notin> set evs4;
@@ -93,14 +93,14 @@
# evs4 \<in> set_mr"
-text{*Note possibility proofs are missing.*}
+text\<open>Note possibility proofs are missing.\<close>
declare Says_imp_knows_Spy [THEN parts.Inj, dest]
declare parts.Body [dest]
declare analz_into_parts [dest]
declare Fake_parts_insert_in_Un [dest]
-text{*General facts about message reception*}
+text\<open>General facts about message reception\<close>
lemma Gets_imp_Says:
"[| Gets B X \<in> set evs; evs \<in> set_mr |] ==> \<exists>A. Says A B X \<in> set evs"
apply (erule rev_mp)
@@ -114,9 +114,9 @@
declare Gets_imp_knows_Spy [THEN parts.Inj, dest]
-subsubsection{*Proofs on keys *}
+subsubsection\<open>Proofs on keys\<close>
-text{*Spy never sees an agent's private keys! (unless it's bad at start)*}
+text\<open>Spy never sees an agent's private keys! (unless it's bad at start)\<close>
lemma Spy_see_private_Key [simp]:
"evs \<in> set_mr
==> (Key(invKey (publicKey b A)) \<in> parts(knows Spy evs)) = (A \<in> bad)"
@@ -144,8 +144,8 @@
qed "signed_keys_in_parts";
???*)
-text{*Proofs on certificates -
- they hold, as in CR, because RCA's keys are secure*}
+text\<open>Proofs on certificates -
+ they hold, as in CR, because RCA's keys are secure\<close>
lemma Crypt_valid_pubEK:
"[| Crypt (priSK RCA) \<lbrace>Agent (CA i), Key EKi, onlyEnc\<rbrace>
@@ -186,20 +186,20 @@
by (blast dest: certificate_valid_pubEK certificate_valid_pubSK)
-text{*Nobody can have used non-existent keys!*}
+text\<open>Nobody can have used non-existent keys!\<close>
lemma new_keys_not_used [rule_format,simp]:
"evs \<in> set_mr
==> Key K \<notin> used evs --> K \<in> symKeys -->
K \<notin> keysFor (parts (knows Spy evs))"
apply (erule set_mr.induct, simp_all)
-apply (force dest!: usedI keysFor_parts_insert) --{*Fake*}
-apply force --{*Message 2*}
-apply (blast dest: Gets_certificate_valid) --{*Message 3*}
-apply force --{*Message 4*}
+apply (force dest!: usedI keysFor_parts_insert) \<comment>\<open>Fake\<close>
+apply force \<comment>\<open>Message 2\<close>
+apply (blast dest: Gets_certificate_valid) \<comment>\<open>Message 3\<close>
+apply force \<comment>\<open>Message 4\<close>
done
-subsubsection{*New Versions: As Above, but Generalized with the Kk Argument*}
+subsubsection\<open>New Versions: As Above, but Generalized with the Kk Argument\<close>
lemma gen_new_keys_not_used [rule_format]:
"evs \<in> set_mr
@@ -232,15 +232,15 @@
K \<in> symKeys; evs \<in> set_mr |] ==> Key K \<in> used evs"
by (blast intro: Crypt_parts_imp_used)
-text{*Rewriting rule for private encryption keys. Analogous rewriting rules
-for other keys aren't needed.*}
+text\<open>Rewriting rule for private encryption keys. Analogous rewriting rules
+for other keys aren't needed.\<close>
lemma parts_image_priEK:
"[|Key (priEK (CA i)) \<in> parts (Key`KK Un (knows Spy evs));
evs \<in> set_mr|] ==> priEK (CA i) \<in> KK | CA i \<in> bad"
by auto
-text{*trivial proof because (priEK (CA i)) never appears even in (parts evs)*}
+text\<open>trivial proof because (priEK (CA i)) never appears even in (parts evs)\<close>
lemma analz_image_priEK:
"evs \<in> set_mr ==>
(Key (priEK (CA i)) \<in> analz (Key`KK Un (knows Spy evs))) =
@@ -248,18 +248,18 @@
by (blast dest!: parts_image_priEK intro: analz_mono [THEN [2] rev_subsetD])
-subsection{*Secrecy of Session Keys*}
+subsection\<open>Secrecy of Session Keys\<close>
-text{*This holds because if (priEK (CA i)) appears in any traffic then it must
- be known to the Spy, by @{text Spy_see_private_Key}*}
+text\<open>This holds because if (priEK (CA i)) appears in any traffic then it must
+ be known to the Spy, by \<open>Spy_see_private_Key\<close>\<close>
lemma merK_neq_priEK:
"[|Key merK \<notin> analz (knows Spy evs);
Key merK \<in> parts (knows Spy evs);
evs \<in> set_mr|] ==> merK \<noteq> priEK C"
by blast
-text{*Lemma for message 4: either merK is compromised (when we don't care)
- or else merK hasn't been used to encrypt K.*}
+text\<open>Lemma for message 4: either merK is compromised (when we don't care)
+ or else merK hasn't been used to encrypt K.\<close>
lemma msg4_priEK_disj:
"[|Gets B \<lbrace>Crypt KM1
(sign K \<lbrace>Agent M, Nonce NM2, Key merSK, Key merEK\<rbrace>),
@@ -292,9 +292,9 @@
add: analz_image_keys_simps abbrev_simps analz_knows_absorb
analz_knows_absorb2 analz_Key_image_insert_eq notin_image_iff
Spy_analz_private_Key analz_image_priEK)
- --{*5 seconds on a 1.6GHz machine*}
-apply spy_analz --{*Fake*}
-apply auto --{*Message 3*}
+ \<comment>\<open>5 seconds on a 1.6GHz machine\<close>
+apply spy_analz \<comment>\<open>Fake\<close>
+apply auto \<comment>\<open>Message 3\<close>
done
lemma symKey_secrecy [rule_format]:
@@ -312,12 +312,12 @@
analz_knows_absorb2 analz_Key_image_insert_eq
symKey_compromise notin_image_iff Spy_analz_private_Key
analz_image_priEK)
-apply spy_analz --{*Fake*}
-apply force --{*Message 1*}
-apply (auto intro: analz_into_parts [THEN usedI] in_parts_Says_imp_used) --{*Message 3*}
+apply spy_analz \<comment>\<open>Fake\<close>
+apply force \<comment>\<open>Message 1\<close>
+apply (auto intro: analz_into_parts [THEN usedI] in_parts_Says_imp_used) \<comment>\<open>Message 3\<close>
done
-subsection{*Unicity *}
+subsection\<open>Unicity\<close>
lemma msg4_Says_imp_Notes:
"[|Says (CA i) M \<lbrace>sign (priSK (CA i)) \<lbrace>Agent M, Nonce NM2, Agent (CA i)\<rbrace>,
@@ -332,8 +332,8 @@
apply (simp_all (no_asm_simp))
done
-text{*Unicity of merSK wrt a given CA:
- merSK uniquely identifies the other components, including merEK*}
+text\<open>Unicity of merSK wrt a given CA:
+ merSK uniquely identifies the other components, including merEK\<close>
lemma merSK_unicity:
"[|Says (CA i) M \<lbrace>sign (priSK(CA i)) \<lbrace>Agent M, Nonce NM2, Agent (CA i)\<rbrace>,
cert M merSK onlySig (priSK (CA i)),
@@ -351,8 +351,8 @@
apply (blast dest!: msg4_Says_imp_Notes)
done
-text{*Unicity of merEK wrt a given CA:
- merEK uniquely identifies the other components, including merSK*}
+text\<open>Unicity of merEK wrt a given CA:
+ merEK uniquely identifies the other components, including merSK\<close>
lemma merEK_unicity:
"[|Says (CA i) M \<lbrace>sign (priSK(CA i)) \<lbrace>Agent M, Nonce NM2, Agent (CA i)\<rbrace>,
cert M merSK onlySig (priSK (CA i)),
@@ -372,19 +372,19 @@
done
-text{* -No interest on secrecy of nonces: they appear to be used
+text\<open>-No interest on secrecy of nonces: they appear to be used
only for freshness.
-No interest on secrecy of merSK or merEK, as in CR.
- -There's no equivalent of the PAN*}
+ -There's no equivalent of the PAN\<close>
-subsection{*Primary Goals of Merchant Registration *}
+subsection\<open>Primary Goals of Merchant Registration\<close>
-subsubsection{*The merchant's certificates really were created by the CA,
-provided the CA is uncompromised *}
+subsubsection\<open>The merchant's certificates really were created by the CA,
+provided the CA is uncompromised\<close>
-text{*The assumption @{term "CA i \<noteq> RCA"} is required: step 2 uses
- certificates of the same form.*}
+text\<open>The assumption @{term "CA i \<noteq> RCA"} is required: step 2 uses
+ certificates of the same form.\<close>
lemma certificate_merSK_valid_lemma [intro]:
"[|Crypt (priSK (CA i)) \<lbrace>Agent M, Key merSK, onlySig\<rbrace>
\<in> parts (knows Spy evs);
@@ -423,7 +423,7 @@
\<lbrace>X, Y, cert M merEK onlyEnc (priSK (CA i)), Z\<rbrace> \<in> set evs"
by auto
-text{*The two certificates - for merSK and for merEK - cannot be proved to
- have originated together*}
+text\<open>The two certificates - for merSK and for merEK - cannot be proved to
+ have originated together\<close>
end
--- a/src/HOL/SET_Protocol/Message_SET.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/SET_Protocol/Message_SET.thy Thu May 26 17:51:22 2016 +0200
@@ -4,29 +4,29 @@
Author: Lawrence C Paulson
*)
-section{*The Message Theory, Modified for SET*}
+section\<open>The Message Theory, Modified for SET\<close>
theory Message_SET
imports Main "~~/src/HOL/Library/Nat_Bijection"
begin
-subsection{*General Lemmas*}
+subsection\<open>General Lemmas\<close>
-text{*Needed occasionally with @{text spy_analz_tac}, e.g. in
- @{text analz_insert_Key_newK}*}
+text\<open>Needed occasionally with \<open>spy_analz_tac\<close>, e.g. in
+ \<open>analz_insert_Key_newK\<close>\<close>
lemma Un_absorb3 [simp] : "A \<union> (B \<union> A) = B \<union> A"
by blast
-text{*Collapses redundant cases in the huge protocol proofs*}
+text\<open>Collapses redundant cases in the huge protocol proofs\<close>
lemmas disj_simps = disj_comms disj_left_absorb disj_assoc
-text{*Effective with assumptions like @{term "K \<notin> range pubK"} and
- @{term "K \<notin> invKey`range pubK"}*}
+text\<open>Effective with assumptions like @{term "K \<notin> range pubK"} and
+ @{term "K \<notin> invKey`range pubK"}\<close>
lemma notin_image_iff: "(y \<notin> f`I) = (\<forall>i\<in>I. f i \<noteq> y)"
by blast
-text{*Effective with the assumption @{term "KK \<subseteq> - (range(invKey o pubK))"} *}
+text\<open>Effective with the assumption @{term "KK \<subseteq> - (range(invKey o pubK))"}\<close>
lemma disjoint_image_iff: "(A <= - (f`I)) = (\<forall>i\<in>I. f i \<notin> A)"
by blast
@@ -35,8 +35,8 @@
type_synonym key = nat
consts
- all_symmetric :: bool --{*true if all keys are symmetric*}
- invKey :: "key=>key" --{*inverse of a symmetric key*}
+ all_symmetric :: bool \<comment>\<open>true if all keys are symmetric\<close>
+ invKey :: "key=>key" \<comment>\<open>inverse of a symmetric key\<close>
specification (invKey)
invKey [simp]: "invKey (invKey K) = K"
@@ -44,27 +44,27 @@
by (rule exI [of _ id], auto)
-text{*The inverse of a symmetric key is itself; that of a public key
- is the private key and vice versa*}
+text\<open>The inverse of a symmetric key is itself; that of a public key
+ is the private key and vice versa\<close>
definition symKeys :: "key set" where
"symKeys == {K. invKey K = K}"
-text{*Agents. We allow any number of certification authorities, cardholders
- merchants, and payment gateways.*}
+text\<open>Agents. We allow any number of certification authorities, cardholders
+ merchants, and payment gateways.\<close>
datatype
agent = CA nat | Cardholder nat | Merchant nat | PG nat | Spy
-text{*Messages*}
+text\<open>Messages\<close>
datatype
- msg = Agent agent --{*Agent names*}
- | Number nat --{*Ordinary integers, timestamps, ...*}
- | Nonce nat --{*Unguessable nonces*}
- | Pan nat --{*Unguessable Primary Account Numbers (??)*}
- | Key key --{*Crypto keys*}
- | Hash msg --{*Hashing*}
- | MPair msg msg --{*Compound messages*}
- | Crypt key msg --{*Encryption, public- or shared-key*}
+ msg = Agent agent \<comment>\<open>Agent names\<close>
+ | Number nat \<comment>\<open>Ordinary integers, timestamps, ...\<close>
+ | Nonce nat \<comment>\<open>Unguessable nonces\<close>
+ | Pan nat \<comment>\<open>Unguessable Primary Account Numbers (??)\<close>
+ | Key key \<comment>\<open>Crypto keys\<close>
+ | Hash msg \<comment>\<open>Hashing\<close>
+ | MPair msg msg \<comment>\<open>Compound messages\<close>
+ | Crypt key msg \<comment>\<open>Encryption, public- or shared-key\<close>
(*Concrete syntax: messages appear as \<open>\<lbrace>A,B,NA\<rbrace>\<close>, etc...*)
@@ -81,9 +81,9 @@
(curry prod_encode 2)
(curry prod_encode 3)
(prod_encode (4,0))"
- --{*maps each agent to a unique natural number, for specifications*}
+ \<comment>\<open>maps each agent to a unique natural number, for specifications\<close>
-text{*The function is indeed injective*}
+text\<open>The function is indeed injective\<close>
lemma inj_nat_of_agent: "inj nat_of_agent"
by (simp add: nat_of_agent_def inj_on_def curry_def prod_encode_eq split: agent.split)
@@ -93,7 +93,7 @@
keysFor :: "msg set => key set"
where "keysFor H = invKey ` {K. \<exists>X. Crypt K X \<in> H}"
-subsubsection{*Inductive definition of all "parts" of a message.*}
+subsubsection\<open>Inductive definition of all "parts" of a message.\<close>
inductive_set
parts :: "msg set => msg set"
@@ -113,7 +113,7 @@
done
-subsubsection{*Inverse of keys*}
+subsubsection\<open>Inverse of keys\<close>
(*Equations hold because constructors are injective; cannot prove for all f*)
lemma Key_image_eq [simp]: "(Key x \<in> Key`A) = (x\<in>A)"
@@ -143,7 +143,7 @@
done
-subsection{*keysFor operator*}
+subsection\<open>keysFor operator\<close>
lemma keysFor_empty [simp]: "keysFor {} = {}"
by (unfold keysFor_def, blast)
@@ -190,7 +190,7 @@
by (unfold keysFor_def, blast)
-subsection{*Inductive relation "parts"*}
+subsection\<open>Inductive relation "parts"\<close>
lemma MPair_parts:
"[| \<lbrace>X,Y\<rbrace> \<in> parts H;
@@ -198,10 +198,10 @@
by (blast dest: parts.Fst parts.Snd)
declare MPair_parts [elim!] parts.Body [dest!]
-text{*NB These two rules are UNSAFE in the formal sense, as they discard the
+text\<open>NB These two rules are UNSAFE in the formal sense, as they discard the
compound message. They work well on THIS FILE.
- @{text MPair_parts} is left as SAFE because it speeds up proofs.
- The Crypt rule is normally kept UNSAFE to avoid breaking up certificates.*}
+ \<open>MPair_parts\<close> is left as SAFE because it speeds up proofs.
+ The Crypt rule is normally kept UNSAFE to avoid breaking up certificates.\<close>
lemma parts_increasing: "H \<subseteq> parts(H)"
by blast
@@ -221,7 +221,7 @@
by (erule parts.induct, fast+)
-subsubsection{*Unions*}
+subsubsection\<open>Unions\<close>
lemma parts_Un_subset1: "parts(G) \<union> parts(H) \<subseteq> parts(G \<union> H)"
by (intro Un_least parts_mono Un_upper1 Un_upper2)
@@ -262,15 +262,15 @@
NOTE: the UN versions are no longer used!*)
-text{*This allows @{text blast} to simplify occurrences of
- @{term "parts(G\<union>H)"} in the assumption.*}
+text\<open>This allows \<open>blast\<close> to simplify occurrences of
+ @{term "parts(G\<union>H)"} in the assumption.\<close>
declare parts_Un [THEN equalityD1, THEN subsetD, THEN UnE, elim!]
lemma parts_insert_subset: "insert X (parts H) \<subseteq> parts(insert X H)"
by (blast intro: parts_mono [THEN [2] rev_subsetD])
-subsubsection{*Idempotence and transitivity*}
+subsubsection\<open>Idempotence and transitivity\<close>
lemma parts_partsD [dest!]: "X\<in> parts (parts H) ==> X\<in> parts H"
by (erule parts.induct, blast+)
@@ -290,7 +290,7 @@
by (force dest!: parts_cut intro: parts_insertI)
-subsubsection{*Rewrite rules for pulling out atomic messages*}
+subsubsection\<open>Rewrite rules for pulling out atomic messages\<close>
lemmas parts_insert_eq_I = equalityI [OF subsetI parts_insert_subset]
@@ -383,11 +383,11 @@
apply (rule_tac x = "N + Suc nat" in exI, auto)
done
-subsection{*Inductive relation "analz"*}
+subsection\<open>Inductive relation "analz"\<close>
-text{*Inductive definition of "analz" -- what can be broken down from a set of
+text\<open>Inductive definition of "analz" -- what can be broken down from a set of
messages, including keys. A form of downward closure. Pairs can
- be taken apart; messages decrypted with known keys.*}
+ be taken apart; messages decrypted with known keys.\<close>
inductive_set
analz :: "msg set => msg set"
@@ -407,7 +407,7 @@
apply (auto dest: Fst Snd)
done
-text{*Making it safe speeds up proofs*}
+text\<open>Making it safe speeds up proofs\<close>
lemma MPair_analz [elim!]:
"[| \<lbrace>X,Y\<rbrace> \<in> analz H;
[| X \<in> analz H; Y \<in> analz H |] ==> P
@@ -440,7 +440,7 @@
lemmas analz_insertI = subset_insertI [THEN analz_mono, THEN [2] rev_subsetD]
-subsubsection{*General equational properties*}
+subsubsection\<open>General equational properties\<close>
lemma analz_empty [simp]: "analz{} = {}"
apply safe
@@ -455,7 +455,7 @@
lemma analz_insert: "insert X (analz H) \<subseteq> analz(insert X H)"
by (blast intro: analz_mono [THEN [2] rev_subsetD])
-subsubsection{*Rewrite rules for pulling out atomic messages*}
+subsubsection\<open>Rewrite rules for pulling out atomic messages\<close>
lemmas analz_insert_eq_I = equalityI [OF subsetI analz_insert]
@@ -568,7 +568,7 @@
done
-subsubsection{*Idempotence and transitivity*}
+subsubsection\<open>Idempotence and transitivity\<close>
lemma analz_analzD [dest!]: "X\<in> analz (analz H) ==> X\<in> analz H"
by (erule analz.induct, blast+)
@@ -594,7 +594,7 @@
by (blast intro: analz_cut analz_insertI)
-text{*A congruence rule for "analz"*}
+text\<open>A congruence rule for "analz"\<close>
lemma analz_subset_cong:
"[| analz G \<subseteq> analz G'; analz H \<subseteq> analz H'
@@ -631,12 +631,12 @@
by (blast intro: analz_UN_analz_lemma analz_mono [THEN [2] rev_subsetD])
-subsection{*Inductive relation "synth"*}
+subsection\<open>Inductive relation "synth"\<close>
-text{*Inductive definition of "synth" -- what can be built up from a set of
+text\<open>Inductive definition of "synth" -- what can be built up from a set of
messages. A form of upward closure. Pairs can be built, messages
encrypted with known keys. Agent names are public domain.
- Numbers can be guessed, but Nonces cannot be.*}
+ Numbers can be guessed, but Nonces cannot be.\<close>
inductive_set
synth :: "msg set => msg set"
@@ -668,7 +668,7 @@
lemma synth_increasing: "H \<subseteq> synth(H)"
by blast
-subsubsection{*Unions*}
+subsubsection\<open>Unions\<close>
(*Converse fails: we can synth more from the union than from the
separate parts, building a compound message using elements of each.*)
@@ -678,7 +678,7 @@
lemma synth_insert: "insert X (synth H) \<subseteq> synth(insert X H)"
by (blast intro: synth_mono [THEN [2] rev_subsetD])
-subsubsection{*Idempotence and transitivity*}
+subsubsection\<open>Idempotence and transitivity\<close>
lemma synth_synthD [dest!]: "X\<in> synth (synth H) ==> X\<in> synth H"
by (erule synth.induct, blast+)
@@ -716,7 +716,7 @@
by (unfold keysFor_def, blast)
-subsubsection{*Combinations of parts, analz and synth*}
+subsubsection\<open>Combinations of parts, analz and synth\<close>
lemma parts_synth [simp]: "parts (synth H) = parts H \<union> synth H"
apply (rule equalityI)
@@ -745,7 +745,7 @@
done
-subsubsection{*For reasoning about the Fake rule in traces*}
+subsubsection\<open>For reasoning about the Fake rule in traces\<close>
lemma parts_insert_subset_Un: "X\<in> G ==> parts(insert X H) \<subseteq> parts G \<union> parts H"
by (rule subset_trans [OF parts_mono parts_Un_subset2], blast)
@@ -806,8 +806,8 @@
declare parts.Body [rule del]
-text{*Rewrites to push in Key and Crypt messages, so that other messages can
- be pulled out using the @{text analz_insert} rules*}
+text\<open>Rewrites to push in Key and Crypt messages, so that other messages can
+ be pulled out using the \<open>analz_insert\<close> rules\<close>
lemmas pushKeys =
insert_commute [of "Key K" "Agent C"]
@@ -828,15 +828,15 @@
insert_commute [of "Crypt X K" "MPair X' Y"]
for X K C N PAN X' Y
-text{*Cannot be added with @{text "[simp]"} -- messages should not always be
- re-ordered.*}
+text\<open>Cannot be added with \<open>[simp]\<close> -- messages should not always be
+ re-ordered.\<close>
lemmas pushes = pushKeys pushCrypts
-subsection{*Tactics useful for many protocol proofs*}
+subsection\<open>Tactics useful for many protocol proofs\<close>
(*<*)
ML
-{*
+\<open>
(*Analysis of Fake cases. Also works for messages that forward unknown parts,
but this application is no longer necessary if analz_insert_eq is used.
DEPENDS UPON "X" REFERRING TO THE FRADULENT MESSAGE *)
@@ -873,7 +873,7 @@
simp_tac ctxt 1,
REPEAT (FIRSTGOAL (resolve_tac ctxt [allI,impI,notI,conjI,iffI])),
DEPTH_SOLVE (atomic_spy_analz_tac ctxt 1)]) i);
-*}
+\<close>
(*>*)
@@ -901,7 +901,7 @@
apply (blast intro: synth_analz_mono [THEN [2] rev_subsetD])
done
-text{*Two generalizations of @{text analz_insert_eq}*}
+text\<open>Two generalizations of \<open>analz_insert_eq\<close>\<close>
lemma gen_analz_insert_eq [rule_format]:
"X \<in> analz H ==> ALL G. H \<subseteq> G --> analz (insert X G) = analz G"
by (blast intro: analz_cut analz_insertI analz_mono [THEN [2] rev_subsetD])
@@ -922,16 +922,16 @@
lemmas Fake_parts_sing_imp_Un = Fake_parts_sing [THEN [2] rev_subsetD]
-method_setup spy_analz = {*
- Scan.succeed (SIMPLE_METHOD' o spy_analz_tac) *}
+method_setup spy_analz = \<open>
+ Scan.succeed (SIMPLE_METHOD' o spy_analz_tac)\<close>
"for proving the Fake case when analz is involved"
-method_setup atomic_spy_analz = {*
- Scan.succeed (SIMPLE_METHOD' o atomic_spy_analz_tac) *}
+method_setup atomic_spy_analz = \<open>
+ Scan.succeed (SIMPLE_METHOD' o atomic_spy_analz_tac)\<close>
"for debugging spy_analz"
-method_setup Fake_insert_simp = {*
- Scan.succeed (SIMPLE_METHOD' o Fake_insert_simp_tac) *}
+method_setup Fake_insert_simp = \<open>
+ Scan.succeed (SIMPLE_METHOD' o Fake_insert_simp_tac)\<close>
"for debugging spy_analz"
end
--- a/src/HOL/SET_Protocol/Public_SET.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/SET_Protocol/Public_SET.thy Thu May 26 17:51:22 2016 +0200
@@ -4,24 +4,24 @@
Author: Lawrence C Paulson
*)
-section{*The Public-Key Theory, Modified for SET*}
+section\<open>The Public-Key Theory, Modified for SET\<close>
theory Public_SET
imports Event_SET
begin
-subsection{*Symmetric and Asymmetric Keys*}
+subsection\<open>Symmetric and Asymmetric Keys\<close>
-text{*definitions influenced by the wish to assign asymmetric keys
+text\<open>definitions influenced by the wish to assign asymmetric keys
- since the beginning - only to RCA and CAs, namely we need a partial
- function on type Agent*}
+ function on type Agent\<close>
-text{*The SET specs mention two signature keys for CAs - we only have one*}
+text\<open>The SET specs mention two signature keys for CAs - we only have one\<close>
consts
publicKey :: "[bool, agent] => key"
- --{*the boolean is TRUE if a signing key*}
+ \<comment>\<open>the boolean is TRUE if a signing key\<close>
abbreviation "pubEK == publicKey False"
abbreviation "pubSK == publicKey True"
@@ -30,8 +30,8 @@
abbreviation "priEK A == invKey (pubEK A)"
abbreviation "priSK A == invKey (pubSK A)"
-text{*By freeness of agents, no two agents have the same key. Since
- @{term "True\<noteq>False"}, no agent has the same signing and encryption keys.*}
+text\<open>By freeness of agents, no two agents have the same key. Since
+ @{term "True\<noteq>False"}, no agent has the same signing and encryption keys.\<close>
specification (publicKey)
injective_publicKey:
@@ -55,15 +55,15 @@
declare privateKey_neq_publicKey [THEN not_sym, iff]
-subsection{*Initial Knowledge*}
+subsection\<open>Initial Knowledge\<close>
-text{*This information is not necessary. Each protocol distributes any needed
+text\<open>This information is not necessary. Each protocol distributes any needed
certificates, and anyway our proofs require a formalization of the Spy's
knowledge only. However, the initial knowledge is as follows:
All agents know RCA's public keys;
RCA and CAs know their own respective keys;
RCA (has already certified and therefore) knows all CAs public keys;
- Spy knows all keys of all bad agents.*}
+ Spy knows all keys of all bad agents.\<close>
overloading initState \<equiv> "initState"
begin
@@ -104,13 +104,13 @@
end
-text{*Injective mapping from agents to PANs: an agent can have only one card*}
+text\<open>Injective mapping from agents to PANs: an agent can have only one card\<close>
consts pan :: "agent => nat"
specification (pan)
inj_pan: "inj pan"
- --{*No two agents have the same PAN*}
+ \<comment>\<open>No two agents have the same PAN\<close>
(*<*)
apply (rule exI [of _ "nat_of_agent"])
apply (simp add: inj_on_def inj_nat_of_agent [THEN inj_eq])
@@ -120,10 +120,10 @@
declare inj_pan [THEN inj_eq, iff]
consts
- XOR :: "nat*nat => nat" --{*no properties are assumed of exclusive-or*}
+ XOR :: "nat*nat => nat" \<comment>\<open>no properties are assumed of exclusive-or\<close>
-subsection{*Signature Primitives*}
+subsection\<open>Signature Primitives\<close>
definition
(* Signature = Message + signed Digest *)
@@ -167,22 +167,22 @@
abbreviation "onlySig == Number (Suc 0)"
abbreviation "authCode == Number (Suc (Suc 0))"
-subsection{*Encryption Primitives*}
+subsection\<open>Encryption Primitives\<close>
definition EXcrypt :: "[key,key,msg,msg] => msg" where
- --{*Extra Encryption*}
+ \<comment>\<open>Extra Encryption\<close>
(*K: the symmetric key EK: the public encryption key*)
"EXcrypt K EK M m =
\<lbrace>Crypt K \<lbrace>M, Hash m\<rbrace>, Crypt EK \<lbrace>Key K, m\<rbrace>\<rbrace>"
definition EXHcrypt :: "[key,key,msg,msg] => msg" where
- --{*Extra Encryption with Hashing*}
+ \<comment>\<open>Extra Encryption with Hashing\<close>
(*K: the symmetric key EK: the public encryption key*)
"EXHcrypt K EK M m =
\<lbrace>Crypt K \<lbrace>M, Hash m\<rbrace>, Crypt EK \<lbrace>Key K, m, Hash M\<rbrace>\<rbrace>"
definition Enc :: "[key,key,key,msg] => msg" where
- --{*Simple Encapsulation with SIGNATURE*}
+ \<comment>\<open>Simple Encapsulation with SIGNATURE\<close>
(*SK: the sender's signing key
K: the symmetric key
EK: the public encryption key*)
@@ -190,12 +190,12 @@
\<lbrace>Crypt K (sign SK M), Crypt EK (Key K)\<rbrace>"
definition EncB :: "[key,key,key,msg,msg] => msg" where
- --{*Encapsulation with Baggage. Keys as above, and baggage b.*}
+ \<comment>\<open>Encapsulation with Baggage. Keys as above, and baggage b.\<close>
"EncB SK K EK M b =
\<lbrace>Enc SK K EK \<lbrace>M, Hash b\<rbrace>, b\<rbrace>"
-subsection{*Basic Properties of pubEK, pubSK, priEK and priSK *}
+subsection\<open>Basic Properties of pubEK, pubSK, priEK and priSK\<close>
lemma publicKey_eq_iff [iff]:
"(publicKey b A = publicKey b' A') = (b=b' & A=A')"
@@ -217,12 +217,12 @@
lemma symKeys_invKey_iff [simp]: "(invKey K \<in> symKeys) = (K \<in> symKeys)"
by (unfold symKeys_def, auto)
-text{*Can be slow (or even loop) as a simprule*}
+text\<open>Can be slow (or even loop) as a simprule\<close>
lemma symKeys_neq_imp_neq: "(K \<in> symKeys) \<noteq> (K' \<in> symKeys) ==> K \<noteq> K'"
by blast
-text{*These alternatives to @{text symKeys_neq_imp_neq} don't seem any better
-in practice.*}
+text\<open>These alternatives to \<open>symKeys_neq_imp_neq\<close> don't seem any better
+in practice.\<close>
lemma publicKey_neq_symKey: "K \<in> symKeys ==> publicKey b A \<noteq> K"
by blast
@@ -241,12 +241,12 @@
by auto
-subsection{*"Image" Equations That Hold for Injective Functions *}
+subsection\<open>"Image" Equations That Hold for Injective Functions\<close>
lemma invKey_image_eq [iff]: "(invKey x \<in> invKey`A) = (x\<in>A)"
by auto
-text{*holds because invKey is injective*}
+text\<open>holds because invKey is injective\<close>
lemma publicKey_image_eq [iff]:
"(publicKey b A \<in> publicKey c ` AS) = (b=c & A\<in>AS)"
by auto
@@ -269,7 +269,7 @@
apply (auto intro: range_eqI)
done
-text{*for proving @{text new_keys_not_used}*}
+text\<open>for proving \<open>new_keys_not_used\<close>\<close>
lemma keysFor_parts_insert:
"[| K \<in> keysFor (parts (insert X H)); X \<in> synth (analz H) |]
==> K \<in> keysFor (parts H) | Key (invKey K) \<in> parts H"
@@ -282,22 +282,22 @@
"[|K \<in> symKeys; Crypt K X \<in> H|] ==> K \<in> keysFor H"
by (drule Crypt_imp_invKey_keysFor, simp)
-text{*Agents see their own private keys!*}
+text\<open>Agents see their own private keys!\<close>
lemma privateKey_in_initStateCA [iff]:
"Key (invKey (publicKey b A)) \<in> initState A"
by (case_tac "A", auto)
-text{*Agents see their own public keys!*}
+text\<open>Agents see their own public keys!\<close>
lemma publicKey_in_initStateCA [iff]: "Key (publicKey b A) \<in> initState A"
by (case_tac "A", auto)
-text{*RCA sees CAs' public keys! *}
+text\<open>RCA sees CAs' public keys!\<close>
lemma pubK_CA_in_initState_RCA [iff]:
"Key (publicKey b (CA i)) \<in> initState RCA"
by auto
-text{*Spy knows all public keys*}
+text\<open>Spy knows all public keys\<close>
lemma knows_Spy_pubEK_i [iff]: "Key (publicKey b A) \<in> knows Spy evs"
apply (induct_tac "evs")
apply (simp_all add: imageI knows_Cons split add: event.split)
@@ -306,13 +306,13 @@
declare knows_Spy_pubEK_i [THEN analz.Inj, iff]
(*needed????*)
-text{*Spy sees private keys of bad agents! [and obviously public keys too]*}
+text\<open>Spy sees private keys of bad agents! [and obviously public keys too]\<close>
lemma knows_Spy_bad_privateKey [intro!]:
"A \<in> bad ==> Key (invKey (publicKey b A)) \<in> knows Spy evs"
by (rule initState_subset_knows [THEN subsetD], simp)
-subsection{*Fresh Nonces for Possibility Theorems*}
+subsection\<open>Fresh Nonces for Possibility Theorems\<close>
lemma Nonce_notin_initState [iff]: "Nonce N \<notin> parts (initState B)"
by (induct_tac "B", auto)
@@ -320,7 +320,7 @@
lemma Nonce_notin_used_empty [simp]: "Nonce N \<notin> used []"
by (simp add: used_Nil)
-text{*In any trace, there is an upper bound N on the greatest nonce in use.*}
+text\<open>In any trace, there is an upper bound N on the greatest nonce in use.\<close>
lemma Nonce_supply_lemma: "\<exists>N. \<forall>n. N<=n --> Nonce n \<notin> used evs"
apply (induct_tac "evs")
apply (rule_tac x = 0 in exI)
@@ -337,10 +337,10 @@
done
-subsection{*Specialized Methods for Possibility Theorems*}
+subsection\<open>Specialized Methods for Possibility Theorems\<close>
ML
-{*
+\<open>
(*Tactic for possibility theorems*)
fun possibility_tac ctxt =
REPEAT (*omit used_Says so that Nonces start from different traces!*)
@@ -356,18 +356,18 @@
(ALLGOALS (asm_simp_tac (ctxt setSolver safe_solver))
THEN
REPEAT_FIRST (resolve_tac ctxt [refl, conjI]))
-*}
+\<close>
-method_setup possibility = {*
- Scan.succeed (SIMPLE_METHOD o possibility_tac) *}
+method_setup possibility = \<open>
+ Scan.succeed (SIMPLE_METHOD o possibility_tac)\<close>
"for proving possibility theorems"
-method_setup basic_possibility = {*
- Scan.succeed (SIMPLE_METHOD o basic_possibility_tac) *}
+method_setup basic_possibility = \<open>
+ Scan.succeed (SIMPLE_METHOD o basic_possibility_tac)\<close>
"for proving possibility theorems"
-subsection{*Specialized Rewriting for Theorems About @{term analz} and Image*}
+subsection\<open>Specialized Rewriting for Theorems About @{term analz} and Image\<close>
lemma insert_Key_singleton: "insert (Key K) H = Key ` {K} Un H"
by blast
@@ -376,17 +376,17 @@
"insert (Key K) (Key`KK Un C) = Key ` (insert K KK) Un C"
by blast
-text{*Needed for @{text DK_fresh_not_KeyCryptKey}*}
+text\<open>Needed for \<open>DK_fresh_not_KeyCryptKey\<close>\<close>
lemma publicKey_in_used [iff]: "Key (publicKey b A) \<in> used evs"
by auto
lemma privateKey_in_used [iff]: "Key (invKey (publicKey b A)) \<in> used evs"
by (blast intro!: initState_into_used)
-text{*Reverse the normal simplification of "image" to build up (not break down)
- the set of keys. Based on @{text analz_image_freshK_ss}, but simpler.*}
+text\<open>Reverse the normal simplification of "image" to build up (not break down)
+ the set of keys. Based on \<open>analz_image_freshK_ss\<close>, but simpler.\<close>
lemmas analz_image_keys_simps =
- simp_thms mem_simps --{*these two allow its use with @{text "only:"}*}
+ simp_thms mem_simps \<comment>\<open>these two allow its use with \<open>only:\<close>\<close>
image_insert [THEN sym] image_Un [THEN sym]
rangeI symKeys_neq_imp_neq
insert_Key_singleton insert_Key_image Un_assoc [THEN sym]
@@ -394,16 +394,16 @@
(*General lemmas proved by Larry*)
-subsection{*Controlled Unfolding of Abbreviations*}
+subsection\<open>Controlled Unfolding of Abbreviations\<close>
-text{*A set is expanded only if a relation is applied to it*}
+text\<open>A set is expanded only if a relation is applied to it\<close>
lemma def_abbrev_simp_relation:
"A = B ==> (A \<in> X) = (B \<in> X) &
(u = A) = (u = B) &
(A = u) = (B = u)"
by auto
-text{*A set is expanded only if one of the given functions is applied to it*}
+text\<open>A set is expanded only if one of the given functions is applied to it\<close>
lemma def_abbrev_simp_function:
"A = B
==> parts (insert A X) = parts (insert B X) &
@@ -411,15 +411,15 @@
keysFor (insert A X) = keysFor (insert B X)"
by auto
-subsubsection{*Special Simplification Rules for @{term signCert}*}
+subsubsection\<open>Special Simplification Rules for @{term signCert}\<close>
-text{*Avoids duplicating X and its components!*}
+text\<open>Avoids duplicating X and its components!\<close>
lemma parts_insert_signCert:
"parts (insert (signCert K X) H) =
insert \<lbrace>X, Crypt K X\<rbrace> (parts (insert (Crypt K X) H))"
by (simp add: signCert_def insert_commute [of X])
-text{*Avoids a case split! [X is always available]*}
+text\<open>Avoids a case split! [X is always available]\<close>
lemma analz_insert_signCert:
"analz (insert (signCert K X) H) =
insert \<lbrace>X, Crypt K X\<rbrace> (insert (Crypt K X) (analz (insert X H)))"
@@ -428,9 +428,9 @@
lemma keysFor_insert_signCert: "keysFor (insert (signCert K X) H) = keysFor H"
by (simp add: signCert_def)
-text{*Controlled rewrite rules for @{term signCert}, just the definitions
+text\<open>Controlled rewrite rules for @{term signCert}, just the definitions
of the others. Encryption primitives are just expanded, despite their huge
- redundancy!*}
+ redundancy!\<close>
lemmas abbrev_simps [simp] =
parts_insert_signCert analz_insert_signCert keysFor_insert_signCert
sign_def [THEN def_abbrev_simp_relation]
@@ -451,7 +451,7 @@
EncB_def [THEN def_abbrev_simp_function]
-subsubsection{*Elimination Rules for Controlled Rewriting *}
+subsubsection\<open>Elimination Rules for Controlled Rewriting\<close>
lemma Enc_partsE:
"!!R. [|Enc SK K EK M \<in> parts H;
@@ -477,7 +477,7 @@
by (unfold EXcrypt_def, blast)
-subsection{*Lemmas to Simplify Expressions Involving @{term analz} *}
+subsection\<open>Lemmas to Simplify Expressions Involving @{term analz}\<close>
lemma analz_knows_absorb:
"Key K \<in> analz (knows Spy evs)
@@ -505,13 +505,13 @@
subset_insertI [THEN [2] subset_trans]
-subsection{*Freshness Lemmas*}
+subsection\<open>Freshness Lemmas\<close>
lemma in_parts_Says_imp_used:
"[|Key K \<in> parts {X}; Says A B X \<in> set evs|] ==> Key K \<in> used evs"
by (blast intro: parts_trans dest!: Says_imp_knows_Spy [THEN parts.Inj])
-text{*A useful rewrite rule with @{term analz_image_keys_simps}*}
+text\<open>A useful rewrite rule with @{term analz_image_keys_simps}\<close>
lemma Crypt_notin_image_Key: "Crypt K X \<notin> Key ` KK"
by auto
--- a/src/HOL/SET_Protocol/Purchase.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/SET_Protocol/Purchase.thy Thu May 26 17:51:22 2016 +0200
@@ -4,17 +4,17 @@
Author: Lawrence C Paulson
*)
-section{*Purchase Phase of SET*}
+section\<open>Purchase Phase of SET\<close>
theory Purchase
imports Public_SET
begin
-text{*
+text\<open>
Note: nonces seem to consist of 20 bytes. That includes both freshness
challenges (Chall-EE, etc.) and important secrets (CardSecret, PANsecret)
-This version omits @{text LID_C} but retains @{text LID_M}. At first glance
+This version omits \<open>LID_C\<close> but retains \<open>LID_M\<close>. At first glance
(Programmer's Guide page 267) it seems that both numbers are just introduced
for the respective convenience of the Cardholder's and Merchant's
system. However, omitting both of them would create a problem of
@@ -25,12 +25,12 @@
bootstrapping message (SET initiation message) which is used by the Merchant
and the Cardholder to agree on the actual transaction. This bootstrapping
message is described in the SET External Interface Guide and ought to generate
-@{text LID_M}. According SET Extern Interface Guide, this number might be a
+\<open>LID_M\<close>. According SET Extern Interface Guide, this number might be a
cookie, an invoice number etc. The Programmer's Guide on page 310, states that
-in absence of @{text LID_M} the protocol must somehow ("outside SET") identify
+in absence of \<open>LID_M\<close> the protocol must somehow ("outside SET") identify
the transaction from OrderDesc, which is assumed to be a searchable text only
field. Thus, it is assumed that the Merchant or the Cardholder somehow agreed
-out-of-bad on the value of @{text LID_M} (for instance a cookie in a web
+out-of-bad on the value of \<open>LID_M\<close> (for instance a cookie in a web
transaction etc.). This out-of-band agreement is expressed with a preliminary
start action in which the merchant and the Cardholder agree on the appropriate
values. Agreed values are stored with a suitable notes action.
@@ -49,40 +49,40 @@
financial network. The data is encrypted by the Cardholder and sent via the
Merchant, such that the data is hidden from the Merchant unless the Acquirer
passes the data back to the Merchant.
---Programmer's Guide, page 271.*}
+--Programmer's Guide, page 271.\<close>
consts
CardSecret :: "nat => nat"
- --{*Maps Cardholders to CardSecrets.
- A CardSecret of 0 means no cerificate, must use unsigned format.*}
+ \<comment>\<open>Maps Cardholders to CardSecrets.
+ A CardSecret of 0 means no cerificate, must use unsigned format.\<close>
PANSecret :: "nat => nat"
- --{*Maps Cardholders to PANSecrets.*}
+ \<comment>\<open>Maps Cardholders to PANSecrets.\<close>
inductive_set
set_pur :: "event list set"
where
- Nil: --{*Initial trace is empty*}
+ Nil: \<comment>\<open>Initial trace is empty\<close>
"[] \<in> set_pur"
-| Fake: --{*The spy MAY say anything he CAN say.*}
+| Fake: \<comment>\<open>The spy MAY say anything he CAN say.\<close>
"[| evsf \<in> set_pur; X \<in> synth(analz(knows Spy evsf)) |]
==> Says Spy B X # evsf \<in> set_pur"
-| Reception: --{*If A sends a message X to B, then B might receive it*}
+| Reception: \<comment>\<open>If A sends a message X to B, then B might receive it\<close>
"[| evsr \<in> set_pur; Says A B X \<in> set evsr |]
==> Gets B X # evsr \<in> set_pur"
| Start:
- --{*Added start event which is out-of-band for SET: the Cardholder and
- the merchant agree on the amounts and uses @{text LID_M} as an
+ \<comment>\<open>Added start event which is out-of-band for SET: the Cardholder and
+ the merchant agree on the amounts and uses \<open>LID_M\<close> as an
identifier.
This is suggested by the External Interface Guide. The Programmer's
- Guide, in absence of @{text LID_M}, states that the merchant uniquely
- identifies the order out of some data contained in OrderDesc.*}
+ Guide, in absence of \<open>LID_M\<close>, states that the merchant uniquely
+ identifies the order out of some data contained in OrderDesc.\<close>
"[|evsStart \<in> set_pur;
Number LID_M \<notin> used evsStart;
C = Cardholder k; M = Merchant i; P = PG j;
@@ -94,7 +94,7 @@
# evsStart \<in> set_pur"
| PInitReq:
- --{*Purchase initialization, page 72 of Formal Protocol Desc.*}
+ \<comment>\<open>Purchase initialization, page 72 of Formal Protocol Desc.\<close>
"[|evsPIReq \<in> set_pur;
Transaction = \<lbrace>Agent M, Agent C, Number OrderDesc, Number PurchAmt\<rbrace>;
Nonce Chall_C \<notin> used evsPIReq;
@@ -103,9 +103,9 @@
==> Says C M \<lbrace>Number LID_M, Nonce Chall_C\<rbrace> # evsPIReq \<in> set_pur"
| PInitRes:
- --{*Merchant replies with his own label XID and the encryption
+ \<comment>\<open>Merchant replies with his own label XID and the encryption
key certificate of his chosen Payment Gateway. Page 74 of Formal
- Protocol Desc. We use @{text LID_M} to identify Cardholder*}
+ Protocol Desc. We use \<open>LID_M\<close> to identify Cardholder\<close>
"[|evsPIRes \<in> set_pur;
Gets M \<lbrace>Number LID_M, Nonce Chall_C\<rbrace> \<in> set evsPIRes;
Transaction = \<lbrace>Agent M, Agent C, Number OrderDesc, Number PurchAmt\<rbrace>;
@@ -121,13 +121,13 @@
# evsPIRes \<in> set_pur"
| PReqUns:
- --{*UNSIGNED Purchase request (CardSecret = 0).
+ \<comment>\<open>UNSIGNED Purchase request (CardSecret = 0).
Page 79 of Formal Protocol Desc.
Merchant never sees the amount in clear. This holds of the real
protocol, where XID identifies the transaction. We omit
\<open>Hash\<lbrace>Number XID, Nonce (CardSecret k)\<rbrace>\<close> from PIHead because
the CardSecret is 0 and because AuthReq treated the unsigned case
- very differently from the signed one anyway.*}
+ very differently from the signed one anyway.\<close>
"!!Chall_C Chall_M OrderDesc P PurchAmt XID evsPReqU.
[|evsPReqU \<in> set_pur;
C = Cardholder k; CardSecret k = 0;
@@ -150,12 +150,12 @@
# evsPReqU \<in> set_pur"
| PReqS:
- --{*SIGNED Purchase request. Page 77 of Formal Protocol Desc.
+ \<comment>\<open>SIGNED Purchase request. Page 77 of Formal Protocol Desc.
We could specify the equation
@{term "PIReqSigned = \<lbrace> PIDualSigned, OIDualSigned \<rbrace>"}, since the
Formal Desc. gives PIHead the same format in the unsigned case.
However, there's little point, as P treats the signed and
- unsigned cases differently.*}
+ unsigned cases differently.\<close>
"!!C Chall_C Chall_M EKj HOD KC2 LID_M M OIData
OIDualSigned OrderDesc P PANData PIData PIDualSigned
PIHead PurchAmt Transaction XID evsPReqS k.
@@ -183,8 +183,8 @@
# Notes C \<lbrace>Key KC2, Agent M\<rbrace>
# evsPReqS \<in> set_pur"
- --{*Authorization Request. Page 92 of Formal Protocol Desc.
- Sent in response to Purchase Request.*}
+ \<comment>\<open>Authorization Request. Page 92 of Formal Protocol Desc.
+ Sent in response to Purchase Request.\<close>
| AuthReq:
"[| evsAReq \<in> set_pur;
Key KM \<notin> used evsAReq; KM \<in> symKeys;
@@ -206,16 +206,16 @@
\<lbrace>Number LID_M, Number XID, Hash OIData, HOD\<rbrace> P_I)
# evsAReq \<in> set_pur"
- --{*Authorization Response has two forms: for UNSIGNED and SIGNED PIs.
+ \<comment>\<open>Authorization Response has two forms: for UNSIGNED and SIGNED PIs.
Page 99 of Formal Protocol Desc.
- PI is a keyword (product!), so we call it @{text P_I}. The hashes HOD and
- HOIData occur independently in @{text P_I} and in M's message.
+ PI is a keyword (product!), so we call it \<open>P_I\<close>. The hashes HOD and
+ HOIData occur independently in \<open>P_I\<close> and in M's message.
The authCode in AuthRes represents the baggage of EncB, which in the
full protocol is [CapToken], [AcqCardMsg], [AuthToken]:
- optional items for split shipments, recurring payments, etc.*}
+ optional items for split shipments, recurring payments, etc.\<close>
| AuthResUns:
- --{*Authorization Response, UNSIGNED*}
+ \<comment>\<open>Authorization Response, UNSIGNED\<close>
"[| evsAResU \<in> set_pur;
C = Cardholder k; M = Merchant i;
Key KP \<notin> used evsAResU; KP \<in> symKeys;
@@ -232,7 +232,7 @@
# evsAResU \<in> set_pur"
| AuthResS:
- --{*Authorization Response, SIGNED*}
+ \<comment>\<open>Authorization Response, SIGNED\<close>
"[| evsAResS \<in> set_pur;
C = Cardholder k;
Key KP \<notin> used evsAResS; KP \<in> symKeys;
@@ -254,7 +254,7 @@
# evsAResS \<in> set_pur"
| PRes:
- --{*Purchase response.*}
+ \<comment>\<open>Purchase response.\<close>
"[| evsPRes \<in> set_pur; KP \<in> symKeys; M = Merchant i;
Transaction = \<lbrace>Agent M, Agent C, Number OrderDesc, Number PurchAmt\<rbrace>;
Gets M (EncB (priSK P) KP (pubEK M)
@@ -279,7 +279,7 @@
inj_CardSecret: "inj CardSecret"
inj_PANSecret: "inj PANSecret"
CardSecret_neq_PANSecret: "CardSecret k \<noteq> PANSecret k'"
- --{*No CardSecret equals any PANSecret*}
+ \<comment>\<open>No CardSecret equals any PANSecret\<close>
apply (rule_tac x="curry prod_encode 0" in exI)
apply (rule_tac x="curry prod_encode 1" in exI)
apply (simp add: prod_encode_eq inj_on_def)
@@ -296,15 +296,15 @@
inj_PANSecret [THEN inj_eq, iff]
-subsection{*Possibility Properties*}
+subsection\<open>Possibility Properties\<close>
lemma Says_to_Gets:
"Says A B X # evs \<in> set_pur ==> Gets B X # Says A B X # evs \<in> set_pur"
by (rule set_pur.Reception, auto)
-text{*Possibility for UNSIGNED purchases. Note that we need to ensure
+text\<open>Possibility for UNSIGNED purchases. Note that we need to ensure
that XID differs from OrderDesc and PurchAmt, since it is supposed to be
-a unique number!*}
+a unique number!\<close>
lemma possibility_Uns:
"[| CardSecret k = 0;
C = Cardholder k; M = Merchant i;
@@ -378,7 +378,7 @@
apply (auto simp add: used_Cons symKeys_neq_imp_neq)
done
-text{*General facts about message reception*}
+text\<open>General facts about message reception\<close>
lemma Gets_imp_Says:
"[| Gets B X \<in> set evs; evs \<in> set_pur |]
==> \<exists>A. Says A B X \<in> set evs"
@@ -392,7 +392,7 @@
declare Gets_imp_knows_Spy [THEN parts.Inj, dest]
-text{*Forwarding lemmas, to aid simplification*}
+text\<open>Forwarding lemmas, to aid simplification\<close>
lemma AuthReq_msg_in_parts_spies:
"[|Gets M \<lbrace>P_I, OIData, HPIData\<rbrace> \<in> set evs;
@@ -405,16 +405,16 @@
by (blast dest: Gets_imp_knows_Spy [THEN analz.Inj])
-subsection{*Proofs on Asymmetric Keys*}
+subsection\<open>Proofs on Asymmetric Keys\<close>
-text{*Private Keys are Secret*}
+text\<open>Private Keys are Secret\<close>
-text{*Spy never sees an agent's private keys! (unless it's bad at start)*}
+text\<open>Spy never sees an agent's private keys! (unless it's bad at start)\<close>
lemma Spy_see_private_Key [simp]:
"evs \<in> set_pur
==> (Key(invKey (publicKey b A)) \<in> parts(knows Spy evs)) = (A \<in> bad)"
apply (erule set_pur.induct)
-apply (frule_tac [9] AuthReq_msg_in_parts_spies) --{*AuthReq*}
+apply (frule_tac [9] AuthReq_msg_in_parts_spies) \<comment>\<open>AuthReq\<close>
apply auto
done
declare Spy_see_private_Key [THEN [2] rev_iffD1, dest!]
@@ -425,14 +425,14 @@
by auto
declare Spy_analz_private_Key [THEN [2] rev_iffD1, dest!]
-text{*rewriting rule for priEK's*}
+text\<open>rewriting rule for priEK's\<close>
lemma parts_image_priEK:
"[|Key (priEK C) \<in> parts (Key`KK Un (knows Spy evs));
evs \<in> set_pur|] ==> priEK C \<in> KK | C \<in> bad"
by auto
-text{*trivial proof because @{term"priEK C"} never appears even in
- @{term "parts evs"}. *}
+text\<open>trivial proof because @{term"priEK C"} never appears even in
+ @{term "parts evs"}.\<close>
lemma analz_image_priEK:
"evs \<in> set_pur ==>
(Key (priEK C) \<in> analz (Key`KK Un (knows Spy evs))) =
@@ -440,7 +440,7 @@
by (blast dest!: parts_image_priEK intro: analz_mono [THEN [2] rev_subsetD])
-subsection{*Public Keys in Certificates are Correct*}
+subsection\<open>Public Keys in Certificates are Correct\<close>
lemma Crypt_valid_pubEK [dest!]:
"[| Crypt (priSK RCA) \<lbrace>Agent C, Key EKi, onlyEnc\<rbrace>
@@ -479,26 +479,26 @@
==> EK = pubEK C"
by (frule Gets_imp_Says, auto)
-method_setup valid_certificate_tac = {*
+method_setup valid_certificate_tac = \<open>
Args.goal_spec >> (fn quant =>
fn ctxt => SIMPLE_METHOD'' quant (fn i =>
EVERY [forward_tac ctxt @{thms Gets_certificate_valid} i,
assume_tac ctxt i, REPEAT (hyp_subst_tac ctxt i)]))
-*}
+\<close>
-subsection{*Proofs on Symmetric Keys*}
+subsection\<open>Proofs on Symmetric Keys\<close>
-text{*Nobody can have used non-existent keys!*}
+text\<open>Nobody can have used non-existent keys!\<close>
lemma new_keys_not_used [rule_format,simp]:
"evs \<in> set_pur
==> Key K \<notin> used evs --> K \<in> symKeys -->
K \<notin> keysFor (parts (knows Spy evs))"
apply (erule set_pur.induct)
-apply (valid_certificate_tac [8]) --{*PReqS*}
-apply (valid_certificate_tac [7]) --{*PReqUns*}
+apply (valid_certificate_tac [8]) \<comment>\<open>PReqS\<close>
+apply (valid_certificate_tac [7]) \<comment>\<open>PReqUns\<close>
apply auto
-apply (force dest!: usedI keysFor_parts_insert) --{*Fake*}
+apply (force dest!: usedI keysFor_parts_insert) \<comment>\<open>Fake\<close>
done
lemma new_keys_not_analzd:
@@ -518,7 +518,7 @@
K \<in> symKeys; evs \<in> set_pur |] ==> Key K \<in> used evs"
by (blast intro: Crypt_parts_imp_used)
-text{*New versions: as above, but generalized to have the KK argument*}
+text\<open>New versions: as above, but generalized to have the KK argument\<close>
lemma gen_new_keys_not_used:
"[|Key K \<notin> used evs; K \<in> symKeys; evs \<in> set_pur |]
@@ -538,7 +538,7 @@
by (simp add: gen_new_keys_not_analzd)
-subsection{*Secrecy of Symmetric Keys*}
+subsection\<open>Secrecy of Symmetric Keys\<close>
lemma Key_analz_image_Key_lemma:
"P --> (Key K \<in> analz (Key`KK Un H)) --> (K\<in>KK | Key K \<in> analz H)
@@ -556,31 +556,31 @@
apply (erule set_pur.induct)
apply (rule_tac [!] allI)+
apply (rule_tac [!] impI [THEN Key_analz_image_Key_lemma, THEN impI])+
-apply (frule_tac [9] AuthReq_msg_in_analz_spies) --{*AReq*}
-apply (valid_certificate_tac [8]) --{*PReqS*}
-apply (valid_certificate_tac [7]) --{*PReqUns*}
+apply (frule_tac [9] AuthReq_msg_in_analz_spies) \<comment>\<open>AReq\<close>
+apply (valid_certificate_tac [8]) \<comment>\<open>PReqS\<close>
+apply (valid_certificate_tac [7]) \<comment>\<open>PReqUns\<close>
apply (simp_all
del: image_insert image_Un imp_disjL
add: analz_image_keys_simps disj_simps
analz_Key_image_insert_eq notin_image_iff
analz_insert_simps analz_image_priEK)
- --{*8 seconds on a 1.6GHz machine*}
-apply spy_analz --{*Fake*}
-apply (blast elim!: ballE)+ --{*PReq: unsigned and signed*}
+ \<comment>\<open>8 seconds on a 1.6GHz machine\<close>
+apply spy_analz \<comment>\<open>Fake\<close>
+apply (blast elim!: ballE)+ \<comment>\<open>PReq: unsigned and signed\<close>
done
-subsection{*Secrecy of Nonces*}
+subsection\<open>Secrecy of Nonces\<close>
-text{*As usual: we express the property as a logical equivalence*}
+text\<open>As usual: we express the property as a logical equivalence\<close>
lemma Nonce_analz_image_Key_lemma:
"P --> (Nonce N \<in> analz (Key`KK Un H)) --> (Nonce N \<in> analz H)
==> P --> (Nonce N \<in> analz (Key`KK Un H)) = (Nonce N \<in> analz H)"
by (blast intro: analz_mono [THEN [2] rev_subsetD])
-text{*The @{text "(no_asm)"} attribute is essential, since it retains
- the quantifier and allows the simprule's condition to itself be simplified.*}
+text\<open>The \<open>(no_asm)\<close> attribute is essential, since it retains
+ the quantifier and allows the simprule's condition to itself be simplified.\<close>
lemma Nonce_compromise [rule_format (no_asm)]:
"evs \<in> set_pur ==>
(\<forall>N KK. (\<forall>K \<in> KK. K \<notin> range(\<lambda>C. priEK C)) -->
@@ -589,17 +589,17 @@
apply (erule set_pur.induct)
apply (rule_tac [!] allI)+
apply (rule_tac [!] impI [THEN Nonce_analz_image_Key_lemma])+
-apply (frule_tac [9] AuthReq_msg_in_analz_spies) --{*AReq*}
-apply (valid_certificate_tac [8]) --{*PReqS*}
-apply (valid_certificate_tac [7]) --{*PReqUns*}
+apply (frule_tac [9] AuthReq_msg_in_analz_spies) \<comment>\<open>AReq\<close>
+apply (valid_certificate_tac [8]) \<comment>\<open>PReqS\<close>
+apply (valid_certificate_tac [7]) \<comment>\<open>PReqUns\<close>
apply (simp_all
del: image_insert image_Un imp_disjL
add: analz_image_keys_simps disj_simps symKey_compromise
analz_Key_image_insert_eq notin_image_iff
analz_insert_simps analz_image_priEK)
- --{*8 seconds on a 1.6GHz machine*}
-apply spy_analz --{*Fake*}
-apply (blast elim!: ballE) --{*PReqS*}
+ \<comment>\<open>8 seconds on a 1.6GHz machine\<close>
+apply spy_analz \<comment>\<open>Fake\<close>
+apply (blast elim!: ballE) \<comment>\<open>PReqS\<close>
done
lemma PANSecret_notin_spies:
@@ -612,39 +612,39 @@
apply (erule rev_mp)
apply (erule set_pur.induct)
apply (frule_tac [9] AuthReq_msg_in_analz_spies)
-apply (valid_certificate_tac [8]) --{*PReqS*}
+apply (valid_certificate_tac [8]) \<comment>\<open>PReqS\<close>
apply (simp_all
del: image_insert image_Un imp_disjL
add: analz_image_keys_simps disj_simps
symKey_compromise pushes sign_def Nonce_compromise
analz_Key_image_insert_eq notin_image_iff
analz_insert_simps analz_image_priEK)
- --{*2.5 seconds on a 1.6GHz machine*}
+ \<comment>\<open>2.5 seconds on a 1.6GHz machine\<close>
apply spy_analz
apply (blast dest!: Gets_imp_knows_Spy [THEN analz.Inj])
apply (blast dest: Says_imp_knows_Spy [THEN analz.Inj]
Gets_imp_knows_Spy [THEN analz.Inj])
-apply (blast dest: Gets_imp_knows_Spy [THEN analz.Inj]) --{*PReqS*}
+apply (blast dest: Gets_imp_knows_Spy [THEN analz.Inj]) \<comment>\<open>PReqS\<close>
apply (blast dest: Says_imp_knows_Spy [THEN analz.Inj]
- Gets_imp_knows_Spy [THEN analz.Inj]) --{*PRes*}
+ Gets_imp_knows_Spy [THEN analz.Inj]) \<comment>\<open>PRes\<close>
done
-text{*This theorem is a bit silly, in that many CardSecrets are 0!
- But then we don't care. NOT USED*}
+text\<open>This theorem is a bit silly, in that many CardSecrets are 0!
+ But then we don't care. NOT USED\<close>
lemma CardSecret_notin_spies:
"evs \<in> set_pur ==> Nonce (CardSecret i) \<notin> parts (knows Spy evs)"
by (erule set_pur.induct, auto)
-subsection{*Confidentiality of PAN*}
+subsection\<open>Confidentiality of PAN\<close>
lemma analz_image_pan_lemma:
"(Pan P \<in> analz (Key`nE Un H)) --> (Pan P \<in> analz H) ==>
(Pan P \<in> analz (Key`nE Un H)) = (Pan P \<in> analz H)"
by (blast intro: analz_mono [THEN [2] rev_subsetD])
-text{*The @{text "(no_asm)"} attribute is essential, since it retains
- the quantifier and allows the simprule's condition to itself be simplified.*}
+text\<open>The \<open>(no_asm)\<close> attribute is essential, since it retains
+ the quantifier and allows the simprule's condition to itself be simplified.\<close>
lemma analz_image_pan [rule_format (no_asm)]:
"evs \<in> set_pur ==>
\<forall>KK. (\<forall>K \<in> KK. K \<notin> range(\<lambda>C. priEK C)) -->
@@ -653,17 +653,17 @@
apply (erule set_pur.induct)
apply (rule_tac [!] allI impI)+
apply (rule_tac [!] analz_image_pan_lemma)+
-apply (frule_tac [9] AuthReq_msg_in_analz_spies) --{*AReq*}
-apply (valid_certificate_tac [8]) --{*PReqS*}
-apply (valid_certificate_tac [7]) --{*PReqUns*}
+apply (frule_tac [9] AuthReq_msg_in_analz_spies) \<comment>\<open>AReq\<close>
+apply (valid_certificate_tac [8]) \<comment>\<open>PReqS\<close>
+apply (valid_certificate_tac [7]) \<comment>\<open>PReqUns\<close>
apply (simp_all
del: image_insert image_Un imp_disjL
add: analz_image_keys_simps
symKey_compromise pushes sign_def
analz_Key_image_insert_eq notin_image_iff
analz_insert_simps analz_image_priEK)
- --{*7 seconds on a 1.6GHz machine*}
-apply spy_analz --{*Fake*}
+ \<comment>\<open>7 seconds on a 1.6GHz machine\<close>
+apply spy_analz \<comment>\<open>Fake\<close>
apply auto
done
@@ -674,7 +674,7 @@
by (simp del: image_insert image_Un
add: analz_image_keys_simps analz_image_pan)
-text{*Confidentiality of the PAN, unsigned case.*}
+text\<open>Confidentiality of the PAN, unsigned case.\<close>
theorem pan_confidentiality_unsigned:
"[| Pan (pan C) \<in> analz(knows Spy evs); C = Cardholder k;
CardSecret k = 0; evs \<in> set_pur|]
@@ -684,21 +684,21 @@
P \<in> bad"
apply (erule rev_mp)
apply (erule set_pur.induct)
-apply (frule_tac [9] AuthReq_msg_in_analz_spies) --{*AReq*}
-apply (valid_certificate_tac [8]) --{*PReqS*}
-apply (valid_certificate_tac [7]) --{*PReqUns*}
+apply (frule_tac [9] AuthReq_msg_in_analz_spies) \<comment>\<open>AReq\<close>
+apply (valid_certificate_tac [8]) \<comment>\<open>PReqS\<close>
+apply (valid_certificate_tac [7]) \<comment>\<open>PReqUns\<close>
apply (simp_all
del: image_insert image_Un imp_disjL
add: analz_image_keys_simps analz_insert_pan analz_image_pan
notin_image_iff
analz_insert_simps analz_image_priEK)
- --{*3 seconds on a 1.6GHz machine*}
-apply spy_analz --{*Fake*}
-apply blast --{*PReqUns: unsigned*}
-apply force --{*PReqS: signed*}
+ \<comment>\<open>3 seconds on a 1.6GHz machine\<close>
+apply spy_analz \<comment>\<open>Fake\<close>
+apply blast \<comment>\<open>PReqUns: unsigned\<close>
+apply force \<comment>\<open>PReqS: signed\<close>
done
-text{*Confidentiality of the PAN, signed case.*}
+text\<open>Confidentiality of the PAN, signed case.\<close>
theorem pan_confidentiality_signed:
"[|Pan (pan C) \<in> analz(knows Spy evs); C = Cardholder k;
CardSecret k \<noteq> 0; evs \<in> set_pur|]
@@ -708,34 +708,34 @@
OIDualSign\<rbrace> \<in> set evs & P \<in> bad"
apply (erule rev_mp)
apply (erule set_pur.induct)
-apply (frule_tac [9] AuthReq_msg_in_analz_spies) --{*AReq*}
-apply (valid_certificate_tac [8]) --{*PReqS*}
-apply (valid_certificate_tac [7]) --{*PReqUns*}
+apply (frule_tac [9] AuthReq_msg_in_analz_spies) \<comment>\<open>AReq\<close>
+apply (valid_certificate_tac [8]) \<comment>\<open>PReqS\<close>
+apply (valid_certificate_tac [7]) \<comment>\<open>PReqUns\<close>
apply (simp_all
del: image_insert image_Un imp_disjL
add: analz_image_keys_simps analz_insert_pan analz_image_pan
notin_image_iff
analz_insert_simps analz_image_priEK)
- --{*3 seconds on a 1.6GHz machine*}
-apply spy_analz --{*Fake*}
-apply force --{*PReqUns: unsigned*}
-apply blast --{*PReqS: signed*}
+ \<comment>\<open>3 seconds on a 1.6GHz machine\<close>
+apply spy_analz \<comment>\<open>Fake\<close>
+apply force \<comment>\<open>PReqUns: unsigned\<close>
+apply blast \<comment>\<open>PReqS: signed\<close>
done
-text{*General goal: that C, M and PG agree on those details of the transaction
+text\<open>General goal: that C, M and PG agree on those details of the transaction
that they are allowed to know about. PG knows about price and account
- details. M knows about the order description and price. C knows both.*}
+ details. M knows about the order description and price. C knows both.\<close>
-subsection{*Proofs Common to Signed and Unsigned Versions*}
+subsection\<open>Proofs Common to Signed and Unsigned Versions\<close>
lemma M_Notes_PG:
"[|Notes M \<lbrace>Number LID_M, Agent P, Agent M, Agent C, etc\<rbrace> \<in> set evs;
evs \<in> set_pur|] ==> \<exists>j. P = PG j"
by (erule rev_mp, erule set_pur.induct, simp_all)
-text{*If we trust M, then @{term LID_M} determines his choice of P
- (Payment Gateway)*}
+text\<open>If we trust M, then @{term LID_M} determines his choice of P
+ (Payment Gateway)\<close>
lemma goodM_gives_correct_PG:
"[| MsgPInitRes =
\<lbrace>Number LID_M, xid, cc, cm, cert P EKj onlyEnc (priSK RCA)\<rbrace>;
@@ -747,7 +747,7 @@
apply clarify
apply (erule rev_mp)
apply (erule set_pur.induct)
-apply (frule_tac [9] AuthReq_msg_in_parts_spies) --{*AuthReq*}
+apply (frule_tac [9] AuthReq_msg_in_parts_spies) \<comment>\<open>AuthReq\<close>
apply simp_all
apply (blast intro: M_Notes_PG)+
done
@@ -762,7 +762,7 @@
EKj = pubEK P"
by (rule refl [THEN goodM_gives_correct_PG, THEN exE], auto)
-text{*When C receives PInitRes, he learns M's choice of P*}
+text\<open>When C receives PInitRes, he learns M's choice of P\<close>
lemma C_verifies_PInitRes:
"[| MsgPInitRes = \<lbrace>Number LID_M, Number XID, Nonce Chall_C, Nonce Chall_M,
cert P EKj onlyEnc (priSK RCA)\<rbrace>;
@@ -775,12 +775,12 @@
apply clarify
apply (erule rev_mp)
apply (erule set_pur.induct)
-apply (frule_tac [9] AuthReq_msg_in_parts_spies) --{*AuthReq*}
+apply (frule_tac [9] AuthReq_msg_in_parts_spies) \<comment>\<open>AuthReq\<close>
apply simp_all
apply (blast intro: M_Notes_PG)+
done
-text{*Corollary of previous one*}
+text\<open>Corollary of previous one\<close>
lemma Says_C_PInitRes:
"[|Says A C (sign (priSK M)
\<lbrace>Number LID_M, Number XID,
@@ -797,8 +797,8 @@
apply (blast dest: refl [THEN C_verifies_PInitRes])
done
-text{*When P receives an AuthReq, he knows that the signed part originated
- with M. PIRes also has a signed message from M....*}
+text\<open>When P receives an AuthReq, he knows that the signed part originated
+ with M. PIRes also has a signed message from M....\<close>
lemma P_verifies_AuthReq:
"[| AuthReqData = \<lbrace>Number LID_M, Number XID, HOIData, HOD\<rbrace>;
Crypt (priSK M) (Hash \<lbrace>AuthReqData, Hash P_I\<rbrace>)
@@ -815,7 +815,7 @@
apply (frule_tac [4] M_Notes_PG, auto)
done
-text{*When M receives AuthRes, he knows that P signed it, including
+text\<open>When M receives AuthRes, he knows that P signed it, including
the identifying tags and the purchase amount, which he can verify.
(Although the spec has SIGNED and UNSIGNED forms of AuthRes, they
send the same message to M.) The conclusion is weak: M is existentially
@@ -823,7 +823,7 @@
the digital envelope weakens the link between @{term MsgAuthRes} and
@{term"priSK M"}. Changing the precondition to refer to
@{term "Crypt K (sign SK M)"} requires assuming @{term K} to be secure, since
- otherwise the Spy could create that message.*}
+ otherwise the Spy could create that message.\<close>
theorem M_verifies_AuthRes:
"[| MsgAuthRes = \<lbrace>\<lbrace>Number LID_M, Number XID, Number PurchAmt\<rbrace>,
Hash authCode\<rbrace>;
@@ -841,16 +841,16 @@
apply clarify
apply (erule rev_mp)
apply (erule set_pur.induct)
-apply (frule_tac [9] AuthReq_msg_in_parts_spies) --{*AuthReq*}
+apply (frule_tac [9] AuthReq_msg_in_parts_spies) \<comment>\<open>AuthReq\<close>
apply simp_all
apply blast+
done
-subsection{*Proofs for Unsigned Purchases*}
+subsection\<open>Proofs for Unsigned Purchases\<close>
-text{*What we can derive from the ASSUMPTION that C issued a purchase request.
- In the unsigned case, we must trust "C": there's no authentication.*}
+text\<open>What we can derive from the ASSUMPTION that C issued a purchase request.
+ In the unsigned case, we must trust "C": there's no authentication.\<close>
lemma C_determines_EKj:
"[| Says C M \<lbrace>EXHcrypt KC1 EKj \<lbrace>PIHead, Hash OIData\<rbrace> (Pan (pan C)),
OIData, Hash\<lbrace>PIHead, Pan (pan C)\<rbrace> \<rbrace> \<in> set evs;
@@ -862,13 +862,13 @@
apply clarify
apply (erule rev_mp)
apply (erule set_pur.induct, simp_all)
-apply (valid_certificate_tac [2]) --{*PReqUns*}
+apply (valid_certificate_tac [2]) \<comment>\<open>PReqUns\<close>
apply auto
apply (blast dest: Gets_imp_Says Says_C_PInitRes)
done
-text{*Unicity of @{term LID_M} between Merchant and Cardholder notes*}
+text\<open>Unicity of @{term LID_M} between Merchant and Cardholder notes\<close>
lemma unique_LID_M:
"[|Notes (Merchant i) \<lbrace>Number LID_M, Agent P, Trans\<rbrace> \<in> set evs;
Notes C \<lbrace>Number LID_M, Agent M, Agent C, Number OD,
@@ -881,7 +881,7 @@
apply (force dest!: Notes_imp_parts_subset_used)
done
-text{*Unicity of @{term LID_M}, for two Merchant Notes events*}
+text\<open>Unicity of @{term LID_M}, for two Merchant Notes events\<close>
lemma unique_LID_M2:
"[|Notes M \<lbrace>Number LID_M, Trans\<rbrace> \<in> set evs;
Notes M \<lbrace>Number LID_M, Trans'\<rbrace> \<in> set evs;
@@ -892,40 +892,40 @@
apply (force dest!: Notes_imp_parts_subset_used)
done
-text{*Lemma needed below: for the case that
- if PRes is present, then @{term LID_M} has been used.*}
+text\<open>Lemma needed below: for the case that
+ if PRes is present, then @{term LID_M} has been used.\<close>
lemma signed_imp_used:
"[| Crypt (priSK M) (Hash X) \<in> parts (knows Spy evs);
M \<notin> bad; evs \<in> set_pur|] ==> parts {X} \<subseteq> used evs"
apply (erule rev_mp)
apply (erule set_pur.induct)
-apply (frule_tac [9] AuthReq_msg_in_parts_spies) --{*AuthReq*}
+apply (frule_tac [9] AuthReq_msg_in_parts_spies) \<comment>\<open>AuthReq\<close>
apply simp_all
apply safe
apply blast+
done
-text{*Similar, with nested Hash*}
+text\<open>Similar, with nested Hash\<close>
lemma signed_Hash_imp_used:
"[| Crypt (priSK C) (Hash \<lbrace>H, Hash X\<rbrace>) \<in> parts (knows Spy evs);
C \<notin> bad; evs \<in> set_pur|] ==> parts {X} \<subseteq> used evs"
apply (erule rev_mp)
apply (erule set_pur.induct)
-apply (frule_tac [9] AuthReq_msg_in_parts_spies) --{*AuthReq*}
+apply (frule_tac [9] AuthReq_msg_in_parts_spies) \<comment>\<open>AuthReq\<close>
apply simp_all
apply safe
apply blast+
done
-text{*Lemma needed below: for the case that
- if PRes is present, then @{text LID_M} has been used.*}
+text\<open>Lemma needed below: for the case that
+ if PRes is present, then \<open>LID_M\<close> has been used.\<close>
lemma PRes_imp_LID_used:
"[| Crypt (priSK M) (Hash \<lbrace>N, X\<rbrace>) \<in> parts (knows Spy evs);
M \<notin> bad; evs \<in> set_pur|] ==> N \<in> used evs"
by (drule signed_imp_used, auto)
-text{*When C receives PRes, he knows that M and P agreed to the purchase details.
- He also knows that P is the same PG as before*}
+text\<open>When C receives PRes, he knows that M and P agreed to the purchase details.
+ He also knows that P is the same PG as before\<close>
lemma C_verifies_PRes_lemma:
"[| Crypt (priSK M) (Hash MsgPRes) \<in> parts (knows Spy evs);
Notes C \<lbrace>Number LID_M, Trans \<rbrace> \<in> set evs;
@@ -945,7 +945,7 @@
apply (erule rev_mp)
apply (erule rev_mp)
apply (erule set_pur.induct)
-apply (frule_tac [9] AuthReq_msg_in_parts_spies) --{*AuthReq*}
+apply (frule_tac [9] AuthReq_msg_in_parts_spies) \<comment>\<open>AuthReq\<close>
apply simp_all
apply blast
apply blast
@@ -954,9 +954,9 @@
apply (blast dest: unique_LID_M)
done
-text{*When the Cardholder receives Purchase Response from an uncompromised
+text\<open>When the Cardholder receives Purchase Response from an uncompromised
Merchant, he knows that M sent it. He also knows that M received a message signed
-by a Payment Gateway chosen by M to authorize the purchase.*}
+by a Payment Gateway chosen by M to authorize the purchase.\<close>
theorem C_verifies_PRes:
"[| MsgPRes = \<lbrace>Number LID_M, Number XID, Nonce Chall_C,
Hash (Number PurchAmt)\<rbrace>;
@@ -974,9 +974,9 @@
apply (auto simp add: sign_def)
done
-subsection{*Proofs for Signed Purchases*}
+subsection\<open>Proofs for Signed Purchases\<close>
-text{*Some Useful Lemmas: the cardholder knows what he is doing*}
+text\<open>Some Useful Lemmas: the cardholder knows what he is doing\<close>
lemma Crypt_imp_Says_Cardholder:
"[| Crypt K \<lbrace>\<lbrace>\<lbrace>Number LID_M, others\<rbrace>, Hash OIData\<rbrace>, Hash PANData\<rbrace>
@@ -994,7 +994,7 @@
apply (erule rev_mp)
apply (erule rev_mp)
apply (erule set_pur.induct, analz_mono_contra)
-apply (frule_tac [9] AuthReq_msg_in_parts_spies) --{*AuthReq*}
+apply (frule_tac [9] AuthReq_msg_in_parts_spies) \<comment>\<open>AuthReq\<close>
apply simp_all
apply auto
done
@@ -1017,16 +1017,16 @@
apply auto
done
-text{*Can't happen: only Merchants create this type of Note*}
+text\<open>Can't happen: only Merchants create this type of Note\<close>
lemma Notes_Cardholder_self_False:
"[|Notes (Cardholder k)
\<lbrace>Number n, Agent P, Agent (Cardholder k), Agent C, etc\<rbrace> \<in> set evs;
evs \<in> set_pur|] ==> False"
by (erule rev_mp, erule set_pur.induct, auto)
-text{*When M sees a dual signature, he knows that it originated with C.
+text\<open>When M sees a dual signature, he knows that it originated with C.
Using XID he knows it was intended for him.
- This guarantee isn't useful to P, who never gets OIData.*}
+ This guarantee isn't useful to P, who never gets OIData.\<close>
theorem M_verifies_Signed_PReq:
"[| MsgDualSign = \<lbrace>HPIData, Hash OIData\<rbrace>;
OIData = \<lbrace>Number LID_M, etc\<rbrace>;
@@ -1041,7 +1041,7 @@
apply (erule rev_mp)
apply (erule rev_mp)
apply (erule set_pur.induct)
-apply (frule_tac [9] AuthReq_msg_in_parts_spies) --{*AuthReq*}
+apply (frule_tac [9] AuthReq_msg_in_parts_spies) \<comment>\<open>AuthReq\<close>
apply simp_all
apply blast
apply (metis subsetD insert_subset parts.Fst parts_increasing signed_Hash_imp_used)
@@ -1049,10 +1049,10 @@
apply (blast dest!: Notes_Cardholder_self_False)
done
-text{*When P sees a dual signature, he knows that it originated with C.
+text\<open>When P sees a dual signature, he knows that it originated with C.
and was intended for M. This guarantee isn't useful to M, who never gets
- PIData. I don't see how to link @{term "PG j"} and @{text LID_M} without
- assuming @{term "M \<notin> bad"}.*}
+ PIData. I don't see how to link @{term "PG j"} and \<open>LID_M\<close> without
+ assuming @{term "M \<notin> bad"}.\<close>
theorem P_verifies_Signed_PReq:
"[| MsgDualSign = \<lbrace>Hash PIData, HOIData\<rbrace>;
PIData = \<lbrace>PIHead, PANData\<rbrace>;
@@ -1102,9 +1102,9 @@
apply (auto simp add: sign_def)
done
-text{*A variant of @{text M_verifies_Signed_PReq} with explicit PI information.
+text\<open>A variant of \<open>M_verifies_Signed_PReq\<close> with explicit PI information.
Even here we cannot be certain about what C sent to M, since a bad
- PG could have replaced the two key fields. (NOT USED)*}
+ PG could have replaced the two key fields. (NOT USED)\<close>
lemma Signed_PReq_imp_Says_Cardholder:
"[| MsgDualSign = \<lbrace>Hash PIData, Hash OIData\<rbrace>;
OIData = \<lbrace>Number LID_M, Number XID, Nonce Chall_C, HOD, etc\<rbrace>;
@@ -1125,7 +1125,7 @@
apply (erule set_pur.induct, simp_all, auto)
done
-text{*When P receives an AuthReq and a dual signature, he knows that C and M
+text\<open>When P receives an AuthReq and a dual signature, he knows that C and M
agree on the essential details. PurchAmt however is never sent by M to
P; instead C and M both send
@{term "HOD = Hash\<lbrace>Number OrderDesc, Number PurchAmt\<rbrace>"}
@@ -1135,7 +1135,7 @@
used in the digital envelopes. On the other hand, M knows the true identity
of PG (namely j'), and sends AReq there; he can't, however, check that
the EXcrypt involves the correct PG's key.
-*}
+\<close>
theorem P_sees_CM_agreement:
"[| AuthReqData = \<lbrace>Number LID_M, Number XID, HOIData, HOD\<rbrace>;
KC \<in> symKeys;
--- a/src/HOL/SMT_Examples/Boogie.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/SMT_Examples/Boogie.thy Thu May 26 17:51:22 2016 +0200
@@ -2,14 +2,14 @@
Author: Sascha Boehme, TU Muenchen
*)
-section {* Proving Boogie-generated verification conditions *}
+section \<open>Proving Boogie-generated verification conditions\<close>
theory Boogie
imports Main
keywords "boogie_file" :: thy_load ("b2i")
begin
-text {*
+text \<open>
HOL-Boogie and its applications are described in:
\begin{itemize}
@@ -22,20 +22,20 @@
Journal of Automated Reasoning, 2009.
\end{itemize}
-*}
+\<close>
-section {* Built-in types and functions of Boogie *}
+section \<open>Built-in types and functions of Boogie\<close>
-subsection {* Integer arithmetics *}
+subsection \<open>Integer arithmetics\<close>
-text {*
-The operations @{text div} and @{text mod} are built-in in Boogie, but
+text \<open>
+The operations \<open>div\<close> and \<open>mod\<close> are built-in in Boogie, but
without a particular semantics due to different interpretations in
programming languages. We assume that each application comes with a
proper axiomatization.
-*}
+\<close>
axiomatization
boogie_div :: "int \<Rightarrow> int \<Rightarrow> int" (infixl "div'_b" 70) and
@@ -43,13 +43,13 @@
-section {* Setup *}
+section \<open>Setup\<close>
ML_file "boogie.ML"
-section {* Verification condition proofs *}
+section \<open>Verification condition proofs\<close>
declare [[smt_oracle = false]]
declare [[smt_read_only_certificates = true]]
--- a/src/HOL/SMT_Examples/SMT_Examples.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/SMT_Examples/SMT_Examples.thy Thu May 26 17:51:22 2016 +0200
@@ -2,7 +2,7 @@
Author: Sascha Boehme, TU Muenchen
*)
-section {* Examples for the SMT binding *}
+section \<open>Examples for the SMT binding\<close>
theory SMT_Examples
imports Complex_Main
@@ -12,7 +12,7 @@
declare [[smt_read_only_certificates = true]]
-section {* Propositional and first-order logic *}
+section \<open>Propositional and first-order logic\<close>
lemma "True" by smt
lemma "p \<or> \<not>p" by smt
@@ -249,9 +249,9 @@
using assms by smt
-section {* Arithmetic *}
+section \<open>Arithmetic\<close>
-subsection {* Linear arithmetic over integers and reals *}
+subsection \<open>Linear arithmetic over integers and reals\<close>
lemma "(3::int) = 3" by smt
lemma "(3::real) = 3" by smt
@@ -287,7 +287,7 @@
(n' = m \<and> m = (n::int))"
by smt
-text{*
+text\<open>
The following example was taken from HOL/ex/PresburgerEx.thy, where it says:
This following theorem proves that all solutions to the
@@ -301,7 +301,7 @@
There, it is proved by "arith". SMT is able to prove this within a fraction
of one second. With proof reconstruction, it takes about 13 seconds on a Core2
processor.
-*}
+\<close>
lemma "\<lbrakk> x3 = \<bar>x2\<bar> - x1; x4 = \<bar>x3\<bar> - x2; x5 = \<bar>x4\<bar> - x3;
x6 = \<bar>x5\<bar> - x4; x7 = \<bar>x6\<bar> - x5; x8 = \<bar>x7\<bar> - x6;
@@ -324,7 +324,7 @@
using assms [[z3_extensions]] by smt
-subsection {* Linear arithmetic with quantifiers *}
+subsection \<open>Linear arithmetic with quantifiers\<close>
lemma "~ (\<exists>x::int. False)" by smt
lemma "~ (\<exists>x::real. False)" by smt
@@ -352,7 +352,7 @@
lemma "\<forall>(a::int) b::int. 0 < b \<or> b < 1" by smt
-subsection {* Non-linear arithmetic over integers and reals *}
+subsection \<open>Non-linear arithmetic over integers and reals\<close>
lemma "a > (0::int) \<Longrightarrow> a*b > 0 \<Longrightarrow> b > 0"
using [[smt_oracle, z3_extensions]]
@@ -378,7 +378,7 @@
using assms by (metis mult_le_0_iff)
-section {* Pairs *}
+section \<open>Pairs\<close>
lemma "fst (x, y) = a \<Longrightarrow> x = a"
using fst_conv by smt
@@ -387,7 +387,7 @@
using fst_conv snd_conv by smt
-section {* Higher-order problems and recursion *}
+section \<open>Higher-order problems and recursion\<close>
lemma "i \<noteq> i1 \<and> i \<noteq> i2 \<Longrightarrow> (f (i1 := v1, i2 := v2)) i = f i"
using fun_upd_same fun_upd_apply by smt
@@ -446,7 +446,7 @@
end
-section {* Monomorphization examples *}
+section \<open>Monomorphization examples\<close>
definition Pred :: "'a \<Rightarrow> bool" where "Pred x = True"
--- a/src/HOL/SMT_Examples/SMT_Tests.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/SMT_Examples/SMT_Tests.thy Thu May 26 17:51:22 2016 +0200
@@ -2,7 +2,7 @@
Author: Sascha Boehme, TU Muenchen
*)
-section {* Tests for the SMT binding *}
+section \<open>Tests for the SMT binding\<close>
theory SMT_Tests
imports Complex_Main
@@ -10,10 +10,10 @@
smt_status
-text {* Most examples are taken from various Isabelle theories and from HOL4. *}
+text \<open>Most examples are taken from various Isabelle theories and from HOL4.\<close>
-section {* Propositional logic *}
+section \<open>Propositional logic\<close>
lemma
"True"
@@ -81,7 +81,7 @@
by smt+
-section {* First-order logic with equality *}
+section \<open>First-order logic with equality\<close>
lemma
"x = x"
@@ -181,7 +181,7 @@
by smt+
-section {* Guidance for quantifier heuristics: patterns *}
+section \<open>Guidance for quantifier heuristics: patterns\<close>
lemma
assumes "\<forall>x.
@@ -198,7 +198,7 @@
using assms by smt
-section {* Meta-logical connectives *}
+section \<open>Meta-logical connectives\<close>
lemma
"True \<Longrightarrow> True"
@@ -222,7 +222,7 @@
by smt+
-section {* Integers *}
+section \<open>Integers\<close>
lemma
"(0::int) = 0"
@@ -377,7 +377,7 @@
by smt+
-section {* Reals *}
+section \<open>Reals\<close>
lemma
"(0::real) = 0"
@@ -487,11 +487,11 @@
by smt+
-section {* Datatypes, Records, and Typedefs *}
+section \<open>Datatypes, Records, and Typedefs\<close>
-subsection {* Without support by the SMT solver *}
+subsection \<open>Without support by the SMT solver\<close>
-subsubsection {* Algebraic datatypes *}
+subsubsection \<open>Algebraic datatypes\<close>
lemma
"x = fst (x, y)"
@@ -529,7 +529,7 @@
by smt+
-subsubsection {* Records *}
+subsubsection \<open>Records\<close>
record point =
cx :: int
@@ -594,7 +594,7 @@
done
-subsubsection {* Type definitions *}
+subsubsection \<open>Type definitions\<close>
typedef int' = "UNIV::int set" by (rule UNIV_witness)
@@ -610,9 +610,9 @@
by (smt n0_def n1_def n2_def plus'_def Abs_int'_inverse Rep_int'_inverse UNIV_I)+
-subsection {* With support by the SMT solver (but without proofs) *}
+subsection \<open>With support by the SMT solver (but without proofs)\<close>
-subsubsection {* Algebraic datatypes *}
+subsubsection \<open>Algebraic datatypes\<close>
lemma
"x = fst (x, y)"
@@ -653,7 +653,7 @@
by smt+
-subsubsection {* Records *}
+subsubsection \<open>Records\<close>
lemma
"\<lparr>cx = x, cy = y\<rparr> = \<lparr>cx = x', cy = y'\<rparr> \<Longrightarrow> x = x' \<and> y = y'"
@@ -716,7 +716,7 @@
by smt
-subsubsection {* Type definitions *}
+subsubsection \<open>Type definitions\<close>
lemma
"n0 \<noteq> n1"
@@ -726,7 +726,7 @@
by (smt n0_def n1_def n2_def plus'_def)+
-section {* Function updates *}
+section \<open>Function updates\<close>
lemma
"(f (i := v)) i = v"
@@ -740,7 +740,7 @@
by smt+
-section {* Sets *}
+section \<open>Sets\<close>
lemma Empty: "x \<notin> {}" by simp
--- a/src/HOL/SMT_Examples/SMT_Word_Examples.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/SMT_Examples/SMT_Word_Examples.thy Thu May 26 17:51:22 2016 +0200
@@ -2,7 +2,7 @@
Author: Sascha Boehme, TU Muenchen
*)
-section {* Word examples for for SMT binding *}
+section \<open>Word examples for for SMT binding\<close>
theory SMT_Word_Examples
imports "~~/src/HOL/Word/Word"
@@ -13,13 +13,13 @@
declare [[smt_certificates = "SMT_Word_Examples.certs"]]
declare [[smt_read_only_certificates = true]]
-text {*
+text \<open>
Currently, there is no proof reconstruction for words.
All lemmas are proved using the oracle mechanism.
-*}
+\<close>
-section {* Bitvector numbers *}
+section \<open>Bitvector numbers\<close>
lemma "(27 :: 4 word) = -5" by smt
lemma "(27 :: 4 word) = 11" by smt
@@ -33,7 +33,7 @@
lemma "x = (5 :: 4 word) \<Longrightarrow> 4 * x = 4" by smt
-section {* Bit-level logic *}
+section \<open>Bit-level logic\<close>
lemma "0b110 AND 0b101 = (0b100 :: 32 word)" by smt
lemma "0b110 OR 0b011 = (0b111 :: 8 word)" by smt
@@ -53,7 +53,7 @@
lemma "w < 256 \<Longrightarrow> (w :: 16 word) AND 0x00FF = w" by smt
-section {* Combined integer-bitvector properties *}
+section \<open>Combined integer-bitvector properties\<close>
lemma
assumes "bv2int 0 = 0"
--- a/src/HOL/SPARK/Examples/Gcd/Greatest_Common_Divisor.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/SPARK/Examples/Gcd/Greatest_Common_Divisor.thy Thu May 26 17:51:22 2016 +0200
@@ -14,20 +14,20 @@
spark_vc procedure_g_c_d_4
proof -
- from `0 < d` have "0 \<le> c mod d" by (rule pos_mod_sign)
- with `0 \<le> c` `0 < d` `c - c sdiv d * d \<noteq> 0` show ?C1
+ from \<open>0 < d\<close> have "0 \<le> c mod d" by (rule pos_mod_sign)
+ with \<open>0 \<le> c\<close> \<open>0 < d\<close> \<open>c - c sdiv d * d \<noteq> 0\<close> show ?C1
by (simp add: sdiv_pos_pos zmod_zdiv_equality')
next
- from `0 \<le> c` `0 < d` `gcd c d = gcd m n` show ?C2
+ from \<open>0 \<le> c\<close> \<open>0 < d\<close> \<open>gcd c d = gcd m n\<close> show ?C2
by (simp add: sdiv_pos_pos zmod_zdiv_equality' gcd_non_0_int)
qed
spark_vc procedure_g_c_d_11
proof -
- from `0 \<le> c` `0 < d` `c - c sdiv d * d = 0`
+ from \<open>0 \<le> c\<close> \<open>0 < d\<close> \<open>c - c sdiv d * d = 0\<close>
have "d dvd c"
by (auto simp add: sdiv_pos_pos dvd_def ac_simps)
- with `0 < d` `gcd c d = gcd m n` show ?C1
+ with \<open>0 < d\<close> \<open>gcd c d = gcd m n\<close> show ?C1
by simp
qed
--- a/src/HOL/SPARK/Examples/Liseq/Longest_Increasing_Subsequence.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/SPARK/Examples/Liseq/Longest_Increasing_Subsequence.thy Thu May 26 17:51:22 2016 +0200
@@ -7,24 +7,24 @@
imports "../../SPARK"
begin
-text {*
+text \<open>
Set of all increasing subsequences in a prefix of an array
-*}
+\<close>
definition iseq :: "(nat \<Rightarrow> 'a::linorder) \<Rightarrow> nat \<Rightarrow> nat set set" where
"iseq xs l = {is. (\<forall>i\<in>is. i < l) \<and>
(\<forall>i\<in>is. \<forall>j\<in>is. i \<le> j \<longrightarrow> xs i \<le> xs j)}"
-text {*
+text \<open>
Length of longest increasing subsequence in a prefix of an array
-*}
+\<close>
definition liseq :: "(nat \<Rightarrow> 'a::linorder) \<Rightarrow> nat \<Rightarrow> nat" where
"liseq xs i = Max (card ` iseq xs i)"
-text {*
+text \<open>
Length of longest increasing subsequence ending at a particular position
-*}
+\<close>
definition liseq' :: "(nat \<Rightarrow> 'a::linorder) \<Rightarrow> nat \<Rightarrow> nat" where
"liseq' xs i = Max (card ` (iseq xs (Suc i) \<inter> {is. Max is = i}))"
@@ -238,8 +238,8 @@
by (simp add: card_insert max_notin)
moreover {
from H j have "xs (Max is) \<le> xs i" by simp
- moreover from `j < i` have "Suc j \<le> i" by simp
- with `is \<in> iseq xs (Suc j)` have "is \<in> iseq xs i"
+ moreover from \<open>j < i\<close> have "Suc j \<le> i" by simp
+ with \<open>is \<in> iseq xs (Suc j)\<close> have "is \<in> iseq xs i"
by (rule iseq_mono)
ultimately have "is \<union> {i} \<in> iseq xs (Suc i)"
by (rule iseq_insert)
@@ -269,52 +269,52 @@
then have "1 < card is" by simp
then have "Max (is - {Max is}) < Max is"
by (rule Max_diff)
- from `is \<in> iseq xs (Suc i)` `1 < card is`
+ from \<open>is \<in> iseq xs (Suc i)\<close> \<open>1 < card is\<close>
have "xs (Max (is - {Max is})) \<le> xs (Max is)"
by (rule iseq_nth)
have "card is - 1 = liseq' xs (Max (is - {i}))"
proof (rule liseq'_eq)
- from `Max is = i` [symmetric] `finite is` `is \<noteq> {}`
+ from \<open>Max is = i\<close> [symmetric] \<open>finite is\<close> \<open>is \<noteq> {}\<close>
show "card is - 1 = card (is - {i})" by simp
next
- from `is \<in> iseq xs (Suc i)` `Max is = i` [symmetric]
+ from \<open>is \<in> iseq xs (Suc i)\<close> \<open>Max is = i\<close> [symmetric]
show "is - {i} \<in> iseq xs (Suc (Max (is - {i})))"
by simp (rule iseq_diff)
next
- from `1 < card is`
+ from \<open>1 < card is\<close>
show "is - {i} \<noteq> {}" by (rule diff_nonempty)
next
fix js
assume "js \<in> iseq xs (Suc (Max (is - {i})))"
"Max js = Max (is - {i})" "finite js" "js \<noteq> {}"
- from `xs (Max (is - {Max is})) \<le> xs (Max is)`
- `Max js = Max (is - {i})` `Max is = i`
+ from \<open>xs (Max (is - {Max is})) \<le> xs (Max is)\<close>
+ \<open>Max js = Max (is - {i})\<close> \<open>Max is = i\<close>
have "xs (Max js) \<le> xs i" by simp
- moreover from `Max is = i` `Max (is - {Max is}) < Max is`
+ moreover from \<open>Max is = i\<close> \<open>Max (is - {Max is}) < Max is\<close>
have "Suc (Max (is - {i})) \<le> i"
by simp
- with `js \<in> iseq xs (Suc (Max (is - {i})))`
+ with \<open>js \<in> iseq xs (Suc (Max (is - {i})))\<close>
have "js \<in> iseq xs i"
by (rule iseq_mono)
ultimately have "js \<union> {i} \<in> iseq xs (Suc i)"
by (rule iseq_insert)
- moreover from `js \<noteq> {}` `finite js` `Max js = Max (is - {i})`
- `Max is = i` [symmetric] `Max (is - {Max is}) < Max is`
+ moreover from \<open>js \<noteq> {}\<close> \<open>finite js\<close> \<open>Max js = Max (is - {i})\<close>
+ \<open>Max is = i\<close> [symmetric] \<open>Max (is - {Max is}) < Max is\<close>
have "Max (js \<union> {i}) = i"
by simp
ultimately have "card (js \<union> {i}) \<le> card is" by (rule R)
- moreover from `Max is = i` [symmetric] `finite js`
- `Max (is - {Max is}) < Max is` `Max js = Max (is - {i})`
+ moreover from \<open>Max is = i\<close> [symmetric] \<open>finite js\<close>
+ \<open>Max (is - {Max is}) < Max is\<close> \<open>Max js = Max (is - {i})\<close>
have "i \<notin> js" by (simp add: max_notin)
- with `finite js`
+ with \<open>finite js\<close>
have "card (js \<union> {i}) = card ((js \<union> {i}) - {i}) + 1"
by simp
ultimately show "card js \<le> card (is - {i})"
- using `i \<notin> js` `Max is = i` [symmetric] `is \<noteq> {}` `finite is`
+ using \<open>i \<notin> js\<close> \<open>Max is = i\<close> [symmetric] \<open>is \<noteq> {}\<close> \<open>finite is\<close>
by simp
qed simp
- with H `Max (is - {Max is}) < Max is`
- `xs (Max (is - {Max is})) \<le> xs (Max is)`
+ with H \<open>Max (is - {Max is}) < Max is\<close>
+ \<open>xs (Max (is - {Max is})) \<le> xs (Max is)\<close>
show ?thesis by auto
qed
qed
@@ -515,7 +515,7 @@
done
-text {* Proof functions *}
+text \<open>Proof functions\<close>
abbreviation (input)
"arr_conv a \<equiv> (\<lambda>n. a (int n))"
@@ -539,7 +539,7 @@
max_ext = "max_ext' :: (int \<Rightarrow> int) \<Rightarrow> int \<Rightarrow> int \<Rightarrow> int"
-text {* The verification conditions *}
+text \<open>The verification conditions\<close>
spark_open "liseq/liseq_length"
--- a/src/HOL/SPARK/Examples/RIPEMD-160/Hash.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/Hash.thy Thu May 26 17:51:22 2016 +0200
@@ -83,10 +83,10 @@
spark_vc function_hash_13
proof -
have loop_suc: "loop__1__i + 2 = (loop__1__i + 1) + 1" by simp
- have "0 <= loop__1__i + 1" using `0 <= loop__1__i` by simp
+ have "0 <= loop__1__i + 1" using \<open>0 <= loop__1__i\<close> by simp
show ?thesis
unfolding loop_suc
- unfolding rounds'_step[OF `0 <= loop__1__i + 1`]
+ unfolding rounds'_step[OF \<open>0 <= loop__1__i + 1\<close>]
unfolding H1[symmetric]
unfolding H18 ..
qed
--- a/src/HOL/SPARK/Examples/RIPEMD-160/R_L.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/R_L.thy Thu May 26 17:51:22 2016 +0200
@@ -12,7 +12,7 @@
spark_vc function_r_l_2
proof -
- from `0 \<le> j` `j \<le> 79`
+ from \<open>0 \<le> j\<close> \<open>j \<le> 79\<close>
show C: ?C1
by (simp add: r_def r_list_def nth_map [symmetric, of _ _ int] del: fun_upd_apply)
(simp add: nth_fun_of_list_eq [of _ _ undefined] del: fun_upd_apply)
@@ -21,7 +21,7 @@
by (simp add: r_list_def)
moreover have "length r_list = 80"
by (simp add: r_list_def)
- ultimately show ?C3 unfolding C using `j \<le> 79`
+ ultimately show ?C3 unfolding C using \<open>j \<le> 79\<close>
by (simp add: r_def list_all_length)
qed
--- a/src/HOL/SPARK/Examples/RIPEMD-160/R_R.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/R_R.thy Thu May 26 17:51:22 2016 +0200
@@ -12,7 +12,7 @@
spark_vc function_r_r_2
proof -
- from `0 \<le> j` `j \<le> 79`
+ from \<open>0 \<le> j\<close> \<open>j \<le> 79\<close>
show C: ?C1
by (simp add: r'_def r'_list_def nth_map [symmetric, of _ _ int] del: fun_upd_apply)
(simp add: nth_fun_of_list_eq [of _ _ undefined] del: fun_upd_apply)
@@ -21,7 +21,7 @@
by (simp add: r'_list_def)
moreover have "length r'_list = 80"
by (simp add: r'_list_def)
- ultimately show ?C3 unfolding C using `j \<le> 79`
+ ultimately show ?C3 unfolding C using \<open>j \<le> 79\<close>
by (simp add: r'_def list_all_length)
qed
--- a/src/HOL/SPARK/Examples/RIPEMD-160/Round.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/Round.thy Thu May 26 17:51:22 2016 +0200
@@ -86,7 +86,7 @@
proof -
have "nat (i + 1) = Suc (nat i)" using assms by simp
show ?thesis
- unfolding `nat (i + 1) = Suc (nat i)` steps_step steps_to_steps'
+ unfolding \<open>nat (i + 1) = Suc (nat i)\<close> steps_step steps_to_steps'
..
qed
@@ -153,19 +153,19 @@
proof -
let ?MM = 4294967296
have AL: "uint(word_of_int e::word32) = e"
- by (rule uint_word_of_int_id[OF `0 <= e` `e <= ?M`])
+ by (rule uint_word_of_int_id[OF \<open>0 <= e\<close> \<open>e <= ?M\<close>])
have CL: "uint(word_of_int b::word32) = b"
- by (rule uint_word_of_int_id[OF `0 <= b` `b <= ?M`])
+ by (rule uint_word_of_int_id[OF \<open>0 <= b\<close> \<open>b <= ?M\<close>])
have DL: "True" ..
have EL: "uint(word_of_int d::word32) = d"
- by (rule uint_word_of_int_id[OF `0 <= d` `d <= ?M`])
+ by (rule uint_word_of_int_id[OF \<open>0 <= d\<close> \<open>d <= ?M\<close>])
have AR: "uint(word_of_int e'::word32) = e'"
- by (rule uint_word_of_int_id[OF `0 <= e'` `e' <= ?M`])
+ by (rule uint_word_of_int_id[OF \<open>0 <= e'\<close> \<open>e' <= ?M\<close>])
have CR: "uint(word_of_int b'::word32) = b'"
- by (rule uint_word_of_int_id[OF `0 <= b'` `b' <= ?M`])
+ by (rule uint_word_of_int_id[OF \<open>0 <= b'\<close> \<open>b' <= ?M\<close>])
have DR: "True" ..
have ER: "uint(word_of_int d'::word32) = d'"
- by (rule uint_word_of_int_id[OF `0 <= d'` `d' <= ?M`])
+ by (rule uint_word_of_int_id[OF \<open>0 <= d'\<close> \<open>d' <= ?M\<close>])
have BL:
"(uint
(word_rotl (s (nat j))
@@ -185,25 +185,25 @@
word_of_int e)"
(is "(uint (word_rotl _ (_ ((((_ + ?F) mod _ + ?X) mod _ + _) mod _))) + _) mod _ = _")
proof -
- have "a mod ?MM = a" using `0 <= a` `a <= ?M`
+ have "a mod ?MM = a" using \<open>0 <= a\<close> \<open>a <= ?M\<close>
by (simp add: int_mod_eq')
- have "?X mod ?MM = ?X" using `0 <= ?X` `?X <= ?M`
+ have "?X mod ?MM = ?X" using \<open>0 <= ?X\<close> \<open>?X <= ?M\<close>
by (simp add: int_mod_eq')
- have "e mod ?MM = e" using `0 <= e` `e <= ?M`
+ have "e mod ?MM = e" using \<open>0 <= e\<close> \<open>e <= ?M\<close>
by (simp add: int_mod_eq')
have "(?MM::int) = 2 ^ len_of TYPE(32)" by simp
show ?thesis
unfolding
word_add_def
- uint_word_of_int_id[OF `0 <= a` `a <= ?M`]
- uint_word_of_int_id[OF `0 <= ?X` `?X <= ?M`]
+ uint_word_of_int_id[OF \<open>0 <= a\<close> \<open>a <= ?M\<close>]
+ uint_word_of_int_id[OF \<open>0 <= ?X\<close> \<open>?X <= ?M\<close>]
int_word_uint
- unfolding `?MM = 2 ^ len_of TYPE(32)`
+ unfolding \<open>?MM = 2 ^ len_of TYPE(32)\<close>
unfolding word_uint.Abs_norm
by (simp add:
- `a mod ?MM = a`
- `e mod ?MM = e`
- `?X mod ?MM = ?X`)
+ \<open>a mod ?MM = a\<close>
+ \<open>e mod ?MM = e\<close>
+ \<open>?X mod ?MM = ?X\<close>)
qed
have BR:
@@ -225,33 +225,33 @@
word_of_int e')"
(is "(uint (word_rotl _ (_ ((((_ + ?F) mod _ + ?X) mod _ + _) mod _))) + _) mod _ = _")
proof -
- have "a' mod ?MM = a'" using `0 <= a'` `a' <= ?M`
+ have "a' mod ?MM = a'" using \<open>0 <= a'\<close> \<open>a' <= ?M\<close>
by (simp add: int_mod_eq')
- have "?X mod ?MM = ?X" using `0 <= ?X` `?X <= ?M`
+ have "?X mod ?MM = ?X" using \<open>0 <= ?X\<close> \<open>?X <= ?M\<close>
by (simp add: int_mod_eq')
- have "e' mod ?MM = e'" using `0 <= e'` `e' <= ?M`
+ have "e' mod ?MM = e'" using \<open>0 <= e'\<close> \<open>e' <= ?M\<close>
by (simp add: int_mod_eq')
have "(?MM::int) = 2 ^ len_of TYPE(32)" by simp
have nat_transfer: "79 - nat j = nat (79 - j)"
- using nat_diff_distrib `0 <= j` `j <= 79`
+ using nat_diff_distrib \<open>0 <= j\<close> \<open>j <= 79\<close>
by simp
show ?thesis
unfolding
word_add_def
- uint_word_of_int_id[OF `0 <= a'` `a' <= ?M`]
- uint_word_of_int_id[OF `0 <= ?X` `?X <= ?M`]
+ uint_word_of_int_id[OF \<open>0 <= a'\<close> \<open>a' <= ?M\<close>]
+ uint_word_of_int_id[OF \<open>0 <= ?X\<close> \<open>?X <= ?M\<close>]
int_word_uint
nat_transfer
- unfolding `?MM = 2 ^ len_of TYPE(32)`
+ unfolding \<open>?MM = 2 ^ len_of TYPE(32)\<close>
unfolding word_uint.Abs_norm
by (simp add:
- `a' mod ?MM = a'`
- `e' mod ?MM = e'`
- `?X mod ?MM = ?X`)
+ \<open>a' mod ?MM = a'\<close>
+ \<open>e' mod ?MM = e'\<close>
+ \<open>?X mod ?MM = ?X\<close>)
qed
show ?thesis
- unfolding steps'_step[OF `0 <= j`] step_hyp[symmetric]
+ unfolding steps'_step[OF \<open>0 <= j\<close>] step_hyp[symmetric]
step_both_def step_r_def step_l_def
by (simp add: AL BL CL DL EL AR BR CR DR ER)
qed
@@ -276,11 +276,11 @@
0 x"
unfolding steps_def
by (simp add:
- uint_word_of_int_id[OF `0 <= ca` `ca <= ?M`]
- uint_word_of_int_id[OF `0 <= cb` `cb <= ?M`]
- uint_word_of_int_id[OF `0 <= cc` `cc <= ?M`]
- uint_word_of_int_id[OF `0 <= cd` `cd <= ?M`]
- uint_word_of_int_id[OF `0 <= ce` `ce <= ?M`])
+ uint_word_of_int_id[OF \<open>0 <= ca\<close> \<open>ca <= ?M\<close>]
+ uint_word_of_int_id[OF \<open>0 <= cb\<close> \<open>cb <= ?M\<close>]
+ uint_word_of_int_id[OF \<open>0 <= cc\<close> \<open>cc <= ?M\<close>]
+ uint_word_of_int_id[OF \<open>0 <= cd\<close> \<open>cd <= ?M\<close>]
+ uint_word_of_int_id[OF \<open>0 <= ce\<close> \<open>ce <= ?M\<close>])
let ?rotate_arg_l =
"((((ca + f 0 cb cc cd) mod 4294967296 +
x (r_l 0)) mod 4294967296 + k_l 0) mod 4294967296)"
@@ -288,35 +288,35 @@
"((((ca + f 79 cb cc cd) mod 4294967296 +
x (r_r 0)) mod 4294967296 + k_r 0) mod 4294967296)"
note returns =
- `wordops__rotate (s_l 0) ?rotate_arg_l =
- rotate_left (s_l 0) ?rotate_arg_l`
- `wordops__rotate (s_r 0) ?rotate_arg_r =
- rotate_left (s_r 0) ?rotate_arg_r`
- `wordops__rotate 10 cc = rotate_left 10 cc`
- `f 0 cb cc cd = f_spec 0 cb cc cd`
- `f 79 cb cc cd = f_spec 79 cb cc cd`
- `k_l 0 = k_l_spec 0`
- `k_r 0 = k_r_spec 0`
- `r_l 0 = r_l_spec 0`
- `r_r 0 = r_r_spec 0`
- `s_l 0 = s_l_spec 0`
- `s_r 0 = s_r_spec 0`
+ \<open>wordops__rotate (s_l 0) ?rotate_arg_l =
+ rotate_left (s_l 0) ?rotate_arg_l\<close>
+ \<open>wordops__rotate (s_r 0) ?rotate_arg_r =
+ rotate_left (s_r 0) ?rotate_arg_r\<close>
+ \<open>wordops__rotate 10 cc = rotate_left 10 cc\<close>
+ \<open>f 0 cb cc cd = f_spec 0 cb cc cd\<close>
+ \<open>f 79 cb cc cd = f_spec 79 cb cc cd\<close>
+ \<open>k_l 0 = k_l_spec 0\<close>
+ \<open>k_r 0 = k_r_spec 0\<close>
+ \<open>r_l 0 = r_l_spec 0\<close>
+ \<open>r_r 0 = r_r_spec 0\<close>
+ \<open>s_l 0 = s_l_spec 0\<close>
+ \<open>s_r 0 = s_r_spec 0\<close>
- note x_borders = `\<forall>i. 0 \<le> i \<and> i \<le> 15 \<longrightarrow> 0 \<le> x i \<and> x i \<le> ?M`
+ note x_borders = \<open>\<forall>i. 0 \<le> i \<and> i \<le> 15 \<longrightarrow> 0 \<le> x i \<and> x i \<le> ?M\<close>
- from `0 <= r_l 0` `r_l 0 <= 15` x_borders
+ from \<open>0 <= r_l 0\<close> \<open>r_l 0 <= 15\<close> x_borders
have "0 \<le> x (r_l 0)" by blast
hence x_lower: "0 <= x (r_l_spec 0)" unfolding returns .
- from `0 <= r_l 0` `r_l 0 <= 15` x_borders
+ from \<open>0 <= r_l 0\<close> \<open>r_l 0 <= 15\<close> x_borders
have "x (r_l 0) <= ?M" by blast
hence x_upper: "x (r_l_spec 0) <= ?M" unfolding returns .
- from `0 <= r_r 0` `r_r 0 <= 15` x_borders
+ from \<open>0 <= r_r 0\<close> \<open>r_r 0 <= 15\<close> x_borders
have "0 \<le> x (r_r 0)" by blast
hence x_lower': "0 <= x (r_r_spec 0)" unfolding returns .
- from `0 <= r_r 0` `r_r 0 <= 15` x_borders
+ from \<open>0 <= r_r 0\<close> \<open>r_r 0 <= 15\<close> x_borders
have "x (r_r 0) <= ?M" by blast
hence x_upper': "x (r_r_spec 0) <= ?M" unfolding returns .
@@ -327,8 +327,8 @@
H2 H4 H6 H8 H10 H2 H4 H6 H8 H10 (* upper bounds *)
H1 H3 H5 H7 H9 H1 H3 H5 H7 H9 (* lower bounds *)
]
- from this[OF x_lower x_upper x_lower' x_upper' `0 <= 0` `0 <= 79`]
- `0 \<le> ca` `0 \<le> ce` x_lower x_lower'
+ from this[OF x_lower x_upper x_lower' x_upper' \<open>0 <= 0\<close> \<open>0 <= 79\<close>]
+ \<open>0 \<le> ca\<close> \<open>0 \<le> ce\<close> x_lower x_lower'
show ?thesis unfolding returns(1) returns(2) unfolding returns
by simp
qed
@@ -347,73 +347,73 @@
have s: "78 - loop__1__j = (79 - (loop__1__j + 1))" by simp
note returns =
- `wordops__rotate (s_l (loop__1__j + 1)) ?rotate_arg_l =
- rotate_left (s_l (loop__1__j + 1)) ?rotate_arg_l`
- `wordops__rotate (s_r (loop__1__j + 1)) ?rotate_arg_r =
- rotate_left (s_r (loop__1__j + 1)) ?rotate_arg_r`
- `f (loop__1__j + 1) clb clc cld =
- f_spec (loop__1__j + 1) clb clc cld`
- `f (78 - loop__1__j) crb crc crd =
- f_spec (78 - loop__1__j) crb crc crd`[simplified s]
- `wordops__rotate 10 clc = rotate_left 10 clc`
- `wordops__rotate 10 crc = rotate_left 10 crc`
- `k_l (loop__1__j + 1) = k_l_spec (loop__1__j + 1)`
- `k_r (loop__1__j + 1) = k_r_spec (loop__1__j + 1)`
- `r_l (loop__1__j + 1) = r_l_spec (loop__1__j + 1)`
- `r_r (loop__1__j + 1) = r_r_spec (loop__1__j + 1)`
- `s_l (loop__1__j + 1) = s_l_spec (loop__1__j + 1)`
- `s_r (loop__1__j + 1) = s_r_spec (loop__1__j + 1)`
+ \<open>wordops__rotate (s_l (loop__1__j + 1)) ?rotate_arg_l =
+ rotate_left (s_l (loop__1__j + 1)) ?rotate_arg_l\<close>
+ \<open>wordops__rotate (s_r (loop__1__j + 1)) ?rotate_arg_r =
+ rotate_left (s_r (loop__1__j + 1)) ?rotate_arg_r\<close>
+ \<open>f (loop__1__j + 1) clb clc cld =
+ f_spec (loop__1__j + 1) clb clc cld\<close>
+ \<open>f (78 - loop__1__j) crb crc crd =
+ f_spec (78 - loop__1__j) crb crc crd\<close>[simplified s]
+ \<open>wordops__rotate 10 clc = rotate_left 10 clc\<close>
+ \<open>wordops__rotate 10 crc = rotate_left 10 crc\<close>
+ \<open>k_l (loop__1__j + 1) = k_l_spec (loop__1__j + 1)\<close>
+ \<open>k_r (loop__1__j + 1) = k_r_spec (loop__1__j + 1)\<close>
+ \<open>r_l (loop__1__j + 1) = r_l_spec (loop__1__j + 1)\<close>
+ \<open>r_r (loop__1__j + 1) = r_r_spec (loop__1__j + 1)\<close>
+ \<open>s_l (loop__1__j + 1) = s_l_spec (loop__1__j + 1)\<close>
+ \<open>s_r (loop__1__j + 1) = s_r_spec (loop__1__j + 1)\<close>
- note x_borders = `\<forall>i. 0 \<le> i \<and> i \<le> 15 \<longrightarrow> 0 \<le> x i \<and> x i \<le> ?M`
+ note x_borders = \<open>\<forall>i. 0 \<le> i \<and> i \<le> 15 \<longrightarrow> 0 \<le> x i \<and> x i \<le> ?M\<close>
- from `0 <= r_l (loop__1__j + 1)` `r_l (loop__1__j + 1) <= 15` x_borders
+ from \<open>0 <= r_l (loop__1__j + 1)\<close> \<open>r_l (loop__1__j + 1) <= 15\<close> x_borders
have "0 \<le> x (r_l (loop__1__j + 1))" by blast
hence x_lower: "0 <= x (r_l_spec (loop__1__j + 1))" unfolding returns .
- from `0 <= r_l (loop__1__j + 1)` `r_l (loop__1__j + 1) <= 15` x_borders
+ from \<open>0 <= r_l (loop__1__j + 1)\<close> \<open>r_l (loop__1__j + 1) <= 15\<close> x_borders
have "x (r_l (loop__1__j + 1)) <= ?M" by blast
hence x_upper: "x (r_l_spec (loop__1__j + 1)) <= ?M" unfolding returns .
- from `0 <= r_r (loop__1__j + 1)` `r_r (loop__1__j + 1) <= 15` x_borders
+ from \<open>0 <= r_r (loop__1__j + 1)\<close> \<open>r_r (loop__1__j + 1) <= 15\<close> x_borders
have "0 \<le> x (r_r (loop__1__j + 1))" by blast
hence x_lower': "0 <= x (r_r_spec (loop__1__j + 1))" unfolding returns .
- from `0 <= r_r (loop__1__j + 1)` `r_r (loop__1__j + 1) <= 15` x_borders
+ from \<open>0 <= r_r (loop__1__j + 1)\<close> \<open>r_r (loop__1__j + 1) <= 15\<close> x_borders
have "x (r_r (loop__1__j + 1)) <= ?M" by blast
hence x_upper': "x (r_r_spec (loop__1__j + 1)) <= ?M" unfolding returns .
- from `0 <= loop__1__j` have "0 <= loop__1__j + 1" by simp
- from `loop__1__j <= 78` have "loop__1__j + 1 <= 79" by simp
+ from \<open>0 <= loop__1__j\<close> have "0 <= loop__1__j + 1" by simp
+ from \<open>loop__1__j <= 78\<close> have "loop__1__j + 1 <= 79" by simp
have "loop__1__j + 1 + 1 = loop__1__j + 2" by simp
note step_from_hyp[OF H1
- `cla <= ?M`
- `clb <= ?M`
- `clc <= ?M`
- `cld <= ?M`
- `cle <= ?M`
- `cra <= ?M`
- `crb <= ?M`
- `crc <= ?M`
- `crd <= ?M`
- `cre <= ?M`
+ \<open>cla <= ?M\<close>
+ \<open>clb <= ?M\<close>
+ \<open>clc <= ?M\<close>
+ \<open>cld <= ?M\<close>
+ \<open>cle <= ?M\<close>
+ \<open>cra <= ?M\<close>
+ \<open>crb <= ?M\<close>
+ \<open>crc <= ?M\<close>
+ \<open>crd <= ?M\<close>
+ \<open>cre <= ?M\<close>
- `0 <= cla`
- `0 <= clb`
- `0 <= clc`
- `0 <= cld`
- `0 <= cle`
- `0 <= cra`
- `0 <= crb`
- `0 <= crc`
- `0 <= crd`
- `0 <= cre`]
+ \<open>0 <= cla\<close>
+ \<open>0 <= clb\<close>
+ \<open>0 <= clc\<close>
+ \<open>0 <= cld\<close>
+ \<open>0 <= cle\<close>
+ \<open>0 <= cra\<close>
+ \<open>0 <= crb\<close>
+ \<open>0 <= crc\<close>
+ \<open>0 <= crd\<close>
+ \<open>0 <= cre\<close>]
from this[OF
x_lower x_upper x_lower' x_upper'
- `0 <= loop__1__j + 1` `loop__1__j + 1 <= 79`]
- `0 \<le> cla` `0 \<le> cle` `0 \<le> cra` `0 \<le> cre` x_lower x_lower'
- show ?thesis unfolding `loop__1__j + 1 + 1 = loop__1__j + 2`
+ \<open>0 <= loop__1__j + 1\<close> \<open>loop__1__j + 1 <= 79\<close>]
+ \<open>0 \<le> cla\<close> \<open>0 \<le> cle\<close> \<open>0 \<le> cra\<close> \<open>0 \<le> cre\<close> x_lower x_lower'
+ show ?thesis unfolding \<open>loop__1__j + 1 + 1 = loop__1__j + 2\<close>
unfolding returns(1) returns(2) unfolding returns
by simp
qed
@@ -437,21 +437,21 @@
x)"
unfolding from_to_id by simp
from
- `0 \<le> ca_init` `ca_init \<le> ?M`
- `0 \<le> cb_init` `cb_init \<le> ?M`
- `0 \<le> cc_init` `cc_init \<le> ?M`
- `0 \<le> cd_init` `cd_init \<le> ?M`
- `0 \<le> ce_init` `ce_init \<le> ?M`
- `0 \<le> cla` `cla \<le> ?M`
- `0 \<le> clb` `clb \<le> ?M`
- `0 \<le> clc` `clc \<le> ?M`
- `0 \<le> cld` `cld \<le> ?M`
- `0 \<le> cle` `cle \<le> ?M`
- `0 \<le> cra` `cra \<le> ?M`
- `0 \<le> crb` `crb \<le> ?M`
- `0 \<le> crc` `crc \<le> ?M`
- `0 \<le> crd` `crd \<le> ?M`
- `0 \<le> cre` `cre \<le> ?M`
+ \<open>0 \<le> ca_init\<close> \<open>ca_init \<le> ?M\<close>
+ \<open>0 \<le> cb_init\<close> \<open>cb_init \<le> ?M\<close>
+ \<open>0 \<le> cc_init\<close> \<open>cc_init \<le> ?M\<close>
+ \<open>0 \<le> cd_init\<close> \<open>cd_init \<le> ?M\<close>
+ \<open>0 \<le> ce_init\<close> \<open>ce_init \<le> ?M\<close>
+ \<open>0 \<le> cla\<close> \<open>cla \<le> ?M\<close>
+ \<open>0 \<le> clb\<close> \<open>clb \<le> ?M\<close>
+ \<open>0 \<le> clc\<close> \<open>clc \<le> ?M\<close>
+ \<open>0 \<le> cld\<close> \<open>cld \<le> ?M\<close>
+ \<open>0 \<le> cle\<close> \<open>cle \<le> ?M\<close>
+ \<open>0 \<le> cra\<close> \<open>cra \<le> ?M\<close>
+ \<open>0 \<le> crb\<close> \<open>crb \<le> ?M\<close>
+ \<open>0 \<le> crc\<close> \<open>crc \<le> ?M\<close>
+ \<open>0 \<le> crd\<close> \<open>crd \<le> ?M\<close>
+ \<open>0 \<le> cre\<close> \<open>cre \<le> ?M\<close>
show ?thesis
unfolding round_def
unfolding steps_to_steps'
--- a/src/HOL/SPARK/Examples/RIPEMD-160/S_L.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/S_L.thy Thu May 26 17:51:22 2016 +0200
@@ -12,7 +12,7 @@
spark_vc function_s_l_2
proof -
- from `0 \<le> j` `j \<le> 79`
+ from \<open>0 \<le> j\<close> \<open>j \<le> 79\<close>
show C: ?C1
by (simp add: s_def s_list_def nth_map [symmetric, of _ _ int] del: fun_upd_apply)
(simp add: nth_fun_of_list_eq [of _ _ undefined] del: fun_upd_apply)
@@ -21,7 +21,7 @@
by (simp add: s_list_def)
moreover have "length s_list = 80"
by (simp add: s_list_def)
- ultimately show ?C3 unfolding C using `j \<le> 79`
+ ultimately show ?C3 unfolding C using \<open>j \<le> 79\<close>
by (simp add: s_def list_all_length)
qed
--- a/src/HOL/SPARK/Examples/RIPEMD-160/S_R.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/S_R.thy Thu May 26 17:51:22 2016 +0200
@@ -12,7 +12,7 @@
spark_vc function_s_r_2
proof -
- from `0 \<le> j` `j \<le> 79`
+ from \<open>0 \<le> j\<close> \<open>j \<le> 79\<close>
show C: ?C1
by (simp add: s'_def s'_list_def nth_map [symmetric, of _ _ int] del: fun_upd_apply)
(simp add: nth_fun_of_list_eq [of _ _ undefined] del: fun_upd_apply)
@@ -21,7 +21,7 @@
by (simp add: s'_list_def)
moreover have "length s'_list = 80"
by (simp add: s'_list_def)
- ultimately show ?C3 unfolding C using `j \<le> 79`
+ ultimately show ?C3 unfolding C using \<open>j \<le> 79\<close>
by (simp add: s'_def list_all_length)
qed
--- a/src/HOL/SPARK/Examples/Sqrt/Sqrt.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/SPARK/Examples/Sqrt/Sqrt.thy Thu May 26 17:51:22 2016 +0200
@@ -11,13 +11,13 @@
spark_vc function_isqrt_4
proof -
- from `0 \<le> r` have "(r = 0 \<or> r = 1 \<or> r = 2) \<or> 2 < r" by auto
+ from \<open>0 \<le> r\<close> have "(r = 0 \<or> r = 1 \<or> r = 2) \<or> 2 < r" by auto
then show "2 * r \<le> 2147483646"
proof
assume "2 < r"
then have "0 < r" by simp
- with `2 < r` have "2 * r < r * r" by (rule mult_strict_right_mono)
- with `r * r \<le> n` and `n \<le> 2147483647` show ?thesis
+ with \<open>2 < r\<close> have "2 * r < r * r" by (rule mult_strict_right_mono)
+ with \<open>r * r \<le> n\<close> and \<open>n \<le> 2147483647\<close> show ?thesis
by simp
qed auto
then show "2 * r \<le> 2147483647" by simp
--- a/src/HOL/SPARK/Manual/Complex_Types.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/SPARK/Manual/Complex_Types.thy Thu May 26 17:51:22 2016 +0200
@@ -39,9 +39,9 @@
spark_vc procedure_initialize_2
proof -
from
- `initialized a loop__1__i`
- `initialized2 (Field1 (a loop__1__i)) 9`
- `initialized3 (Field1 (a loop__1__i)) 9 (pos Sun)`
+ \<open>initialized a loop__1__i\<close>
+ \<open>initialized2 (Field1 (a loop__1__i)) 9\<close>
+ \<open>initialized3 (Field1 (a loop__1__i)) 9 (pos Sun)\<close>
show ?C1
apply (auto simp add: initialized_def initialized2_def initialized3_def)
apply (case_tac "j = 9")
@@ -58,11 +58,11 @@
spark_vc procedure_initialize_4
proof -
- from `initialized a loop__1__i`
+ from \<open>initialized a loop__1__i\<close>
show ?C1 by (simp add: initialized_def)
from
- `initialized2 (Field1 (a loop__1__i)) loop__2__j`
- `initialized3 (Field1 (a loop__1__i)) loop__2__j (pos Sun)`
+ \<open>initialized2 (Field1 (a loop__1__i)) loop__2__j\<close>
+ \<open>initialized3 (Field1 (a loop__1__i)) loop__2__j (pos Sun)\<close>
show ?C2
apply (auto simp add: initialized2_def initialized3_def)
apply (case_tac "j = loop__2__j")
@@ -79,13 +79,13 @@
spark_vc procedure_initialize_6
proof -
- from `initialized a loop__1__i`
+ from \<open>initialized a loop__1__i\<close>
show ?C1 by (simp add: initialized_def)
- from `initialized2 (Field1 (a loop__1__i)) loop__2__j`
+ from \<open>initialized2 (Field1 (a loop__1__i)) loop__2__j\<close>
show ?C2 by (simp add: initialized2_def initialized3_def)
from
- `initialized3 (Field1 (a loop__1__i)) loop__2__j (pos loop__3__k)`
- `loop__3__k < Sun`
+ \<open>initialized3 (Field1 (a loop__1__i)) loop__2__j (pos loop__3__k)\<close>
+ \<open>loop__3__k < Sun\<close>
rangeI [of pos loop__3__k]
show ?C3
apply (auto simp add: initialized3_def succ_def less_pos pos_val range_pos)
@@ -98,9 +98,9 @@
spark_vc procedure_initialize_9
using
- `initialized a 9`
- `initialized2 (Field1 (a 9)) 9`
- `initialized3 (Field1 (a 9)) 9 (pos Sun)`
+ \<open>initialized a 9\<close>
+ \<open>initialized2 (Field1 (a 9)) 9\<close>
+ \<open>initialized3 (Field1 (a 9)) 9 (pos Sun)\<close>
apply (auto simp add: initialized_def initialized2_def initialized3_def)
apply (case_tac "j = 9")
apply auto
--- a/src/HOL/SPARK/Manual/Example_Verification.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/SPARK/Manual/Example_Verification.thy Thu May 26 17:51:22 2016 +0200
@@ -4,9 +4,9 @@
begin
(*>*)
-chapter {* Verifying an Example Program *}
+chapter \<open>Verifying an Example Program\<close>
-text {*
+text \<open>
\label{sec:example-verification}
\begin{figure}
\lstinputlisting{Gcd.ads}
@@ -24,11 +24,11 @@
the correctness of an example program. As an example, we use a program for computing
the \emph{greatest common divisor} of two natural numbers shown in \figref{fig:gcd-prog},
which has been taken from the book about \SPARK{} by Barnes @{cite \<open>\S 11.6\<close> Barnes}.
-*}
+\<close>
-section {* Importing \SPARK{} VCs into Isabelle *}
+section \<open>Importing \SPARK{} VCs into Isabelle\<close>
-text {*
+text \<open>
In order to specify that the \SPARK{} procedure \texttt{G\_C\_D} behaves like its
mathematical counterpart, Barnes introduces a \emph{proof function} \texttt{Gcd}
in the package specification. Invoking the \SPARK{} Examiner and Simplifier on
@@ -76,23 +76,23 @@
Context: \\
\ \\
\begin{tabular}{ll}
-fixes & @{text "m ::"}\ "@{text int}" \\
-and & @{text "n ::"}\ "@{text int}" \\
-and & @{text "c ::"}\ "@{text int}" \\
-and & @{text "d ::"}\ "@{text int}" \\
-assumes & @{text "g_c_d_rules1:"}\ "@{text "0 \<le> integer__size"}" \\
-and & @{text "g_c_d_rules6:"}\ "@{text "0 \<le> natural__size"}" \\
+fixes & \<open>m ::\<close>\ "\<open>int\<close>" \\
+and & \<open>n ::\<close>\ "\<open>int\<close>" \\
+and & \<open>c ::\<close>\ "\<open>int\<close>" \\
+and & \<open>d ::\<close>\ "\<open>int\<close>" \\
+assumes & \<open>g_c_d_rules1:\<close>\ "\<open>0 \<le> integer__size\<close>" \\
+and & \<open>g_c_d_rules6:\<close>\ "\<open>0 \<le> natural__size\<close>" \\
\multicolumn{2}{l}{notes definition} \\
-\multicolumn{2}{l}{\hspace{2ex}@{text "defns ="}\ `@{text "integer__first = - 2147483648"}`} \\
-\multicolumn{2}{l}{\hspace{4ex}`@{text "integer__last = 2147483647"}`} \\
+\multicolumn{2}{l}{\hspace{2ex}\<open>defns =\<close>\ `\<open>integer__first = - 2147483648\<close>`} \\
+\multicolumn{2}{l}{\hspace{4ex}`\<open>integer__last = 2147483647\<close>`} \\
\multicolumn{2}{l}{\hspace{4ex}\dots}
\end{tabular}\ \\[1.5ex]
\ \\
Definitions: \\
\ \\
\begin{tabular}{ll}
-@{text "g_c_d_rules2:"} & @{text "integer__first = - 2147483648"} \\
-@{text "g_c_d_rules3:"} & @{text "integer__last = 2147483647"} \\
+\<open>g_c_d_rules2:\<close> & \<open>integer__first = - 2147483648\<close> \\
+\<open>g_c_d_rules3:\<close> & \<open>integer__last = 2147483647\<close> \\
\dots
\end{tabular}\ \\[1.5ex]
\ \\
@@ -100,25 +100,25 @@
\ \\
path(s) from assertion of line 10 to assertion of line 10 \\
\ \\
-@{text procedure_g_c_d_4}\ (unproved) \\
+\<open>procedure_g_c_d_4\<close>\ (unproved) \\
\ \ \begin{tabular}{ll}
-assumes & @{text "H1:"}\ "@{text "0 \<le> c"}" \\
-and & @{text "H2:"}\ "@{text "0 < d"}" \\
-and & @{text "H3:"}\ "@{text "gcd c d = gcd m n"}" \\
+assumes & \<open>H1:\<close>\ "\<open>0 \<le> c\<close>" \\
+and & \<open>H2:\<close>\ "\<open>0 < d\<close>" \\
+and & \<open>H3:\<close>\ "\<open>gcd c d = gcd m n\<close>" \\
\dots \\
-shows & "@{text "0 < c - c sdiv d * d"}" \\
-and & "@{text "gcd d (c - c sdiv d * d) = gcd m n"}
+shows & "\<open>0 < c - c sdiv d * d\<close>" \\
+and & "\<open>gcd d (c - c sdiv d * d) = gcd m n\<close>
\end{tabular}\ \\[1.5ex]
\ \\
path(s) from assertion of line 10 to finish \\
\ \\
-@{text procedure_g_c_d_11}\ (unproved) \\
+\<open>procedure_g_c_d_11\<close>\ (unproved) \\
\ \ \begin{tabular}{ll}
-assumes & @{text "H1:"}\ "@{text "0 \<le> c"}" \\
-and & @{text "H2:"}\ "@{text "0 < d"}" \\
-and & @{text "H3:"}\ "@{text "gcd c d = gcd m n"}" \\
+assumes & \<open>H1:\<close>\ "\<open>0 \<le> c\<close>" \\
+and & \<open>H2:\<close>\ "\<open>0 < d\<close>" \\
+and & \<open>H3:\<close>\ "\<open>gcd c d = gcd m n\<close>" \\
\dots \\
-shows & "@{text "d = gcd m n"}"
+shows & "\<open>d = gcd m n\<close>"
\end{tabular}
\end{flushleft}
\caption{Output of \isa{\isacommand{spark\_status}} for \texttt{g\_c\_d.siv}}
@@ -139,81 +139,81 @@
FDL rules of the form ``\dots\ \texttt{may\_be\_replaced\_by}\ \dots'' are
turned into native Isabelle definitions, whereas other rules are modelled
as assumptions.
-*}
+\<close>
-section {* Proving the VCs *}
+section \<open>Proving the VCs\<close>
-text {*
+text \<open>
\label{sec:proving-vcs}
-The two open VCs are @{text procedure_g_c_d_4} and @{text procedure_g_c_d_11},
-both of which contain the @{text gcd} proof function that the \SPARK{} Simplifier
+The two open VCs are \<open>procedure_g_c_d_4\<close> and \<open>procedure_g_c_d_11\<close>,
+both of which contain the \<open>gcd\<close> proof function that the \SPARK{} Simplifier
does not know anything about. The proof of a particular VC can be started with
the \isa{\isacommand{spark\_vc}} command, which is similar to the standard
\isa{\isacommand{lemma}} and \isa{\isacommand{theorem}} commands, with the
difference that it only takes a name of a VC but no formula as an argument.
A VC can have several conclusions that can be referenced by the identifiers
-@{text "?C1"}, @{text "?C2"}, etc. If there is just one conclusion, it can
-also be referenced by @{text "?thesis"}. It is important to note that the
-\texttt{div} operator of FDL behaves differently from the @{text div} operator
+\<open>?C1\<close>, \<open>?C2\<close>, etc. If there is just one conclusion, it can
+also be referenced by \<open>?thesis\<close>. It is important to note that the
+\texttt{div} operator of FDL behaves differently from the \<open>div\<close> operator
of Isabelle/HOL on negative numbers. The former always truncates towards zero,
whereas the latter truncates towards minus infinity. This is why the FDL
-\texttt{div} operator is mapped to the @{text sdiv} operator in Isabelle/HOL,
+\texttt{div} operator is mapped to the \<open>sdiv\<close> operator in Isabelle/HOL,
which is defined as
@{thm [display] sdiv_def}
For example, we have that
@{lemma "-5 sdiv 4 = -1" by (simp add: sdiv_neg_pos)}, but
@{lemma "(-5::int) div 4 = -2" by simp}.
-For non-negative dividend and divisor, @{text sdiv} is equivalent to @{text div},
-as witnessed by theorem @{text sdiv_pos_pos}:
+For non-negative dividend and divisor, \<open>sdiv\<close> is equivalent to \<open>div\<close>,
+as witnessed by theorem \<open>sdiv_pos_pos\<close>:
@{thm [display,mode=no_brackets] sdiv_pos_pos}
In contrast, the behaviour of the FDL \texttt{mod} operator is equivalent to
the one of Isabelle/HOL. Moreover, since FDL has no counterpart of the \SPARK{}
operator \textbf{rem}, the \SPARK{} expression \texttt{c}\ \textbf{rem}\ \texttt{d}
-just becomes @{text "c - c sdiv d * d"} in Isabelle. The first conclusion of
-@{text procedure_g_c_d_4} requires us to prove that the remainder of @{text c}
-and @{text d} is greater than @{text 0}. To do this, we use the theorem
-@{text zmod_zdiv_equality'} describing the correspondence between @{text div}
-and @{text mod}
+just becomes \<open>c - c sdiv d * d\<close> in Isabelle. The first conclusion of
+\<open>procedure_g_c_d_4\<close> requires us to prove that the remainder of \<open>c\<close>
+and \<open>d\<close> is greater than \<open>0\<close>. To do this, we use the theorem
+\<open>zmod_zdiv_equality'\<close> describing the correspondence between \<open>div\<close>
+and \<open>mod\<close>
@{thm [display] zmod_zdiv_equality'}
-together with the theorem @{text pos_mod_sign} saying that the result of the
-@{text mod} operator is non-negative when applied to a non-negative divisor:
+together with the theorem \<open>pos_mod_sign\<close> saying that the result of the
+\<open>mod\<close> operator is non-negative when applied to a non-negative divisor:
@{thm [display] pos_mod_sign}
-We will also need the aforementioned theorem @{text sdiv_pos_pos} in order for
-the standard Isabelle/HOL theorems about @{text div} to be applicable
-to the VC, which is formulated using @{text sdiv} rather that @{text div}.
-Note that the proof uses \texttt{`@{text "0 \<le> c"}`} and \texttt{`@{text "0 < d"}`}
-rather than @{text H1} and @{text H2} to refer to the hypotheses of the current
+We will also need the aforementioned theorem \<open>sdiv_pos_pos\<close> in order for
+the standard Isabelle/HOL theorems about \<open>div\<close> to be applicable
+to the VC, which is formulated using \<open>sdiv\<close> rather that \<open>div\<close>.
+Note that the proof uses \texttt{`\<open>0 \<le> c\<close>`} and \texttt{`\<open>0 < d\<close>`}
+rather than \<open>H1\<close> and \<open>H2\<close> to refer to the hypotheses of the current
VC. While the latter variant seems more compact, it is not particularly robust,
since the numbering of hypotheses can easily change if the corresponding
program is modified, making the proof script hard to adjust when there are many hypotheses.
-Moreover, proof scripts using abbreviations like @{text H1} and @{text H2}
+Moreover, proof scripts using abbreviations like \<open>H1\<close> and \<open>H2\<close>
are hard to read without assistance from Isabelle.
-The second conclusion of @{text procedure_g_c_d_4} requires us to prove that
-the @{text gcd} of @{text d} and the remainder of @{text c} and @{text d}
-is equal to the @{text gcd} of the original input values @{text m} and @{text n},
+The second conclusion of \<open>procedure_g_c_d_4\<close> requires us to prove that
+the \<open>gcd\<close> of \<open>d\<close> and the remainder of \<open>c\<close> and \<open>d\<close>
+is equal to the \<open>gcd\<close> of the original input values \<open>m\<close> and \<open>n\<close>,
which is the actual \emph{invariant} of the procedure. This is a consequence
-of theorem @{text gcd_non_0_int}
+of theorem \<open>gcd_non_0_int\<close>
@{thm [display] gcd_non_0_int}
-Again, we also need theorems @{text zmod_zdiv_equality'} and @{text sdiv_pos_pos}
+Again, we also need theorems \<open>zmod_zdiv_equality'\<close> and \<open>sdiv_pos_pos\<close>
to justify that \SPARK{}'s \textbf{rem} operator is equivalent to Isabelle's
-@{text mod} operator for non-negative operands.
-The VC @{text procedure_g_c_d_11} says that if the loop invariant holds before
+\<open>mod\<close> operator for non-negative operands.
+The VC \<open>procedure_g_c_d_11\<close> says that if the loop invariant holds before
the last iteration of the loop, the postcondition of the procedure will hold
after execution of the loop body. To prove this, we observe that the remainder
-of @{text c} and @{text d}, and hence @{text "c mod d"} is @{text 0} when exiting
-the loop. This implies that @{text "gcd c d = d"}, since @{text c} is divisible
-by @{text d}, so the conclusion follows using the assumption @{text "gcd c d = gcd m n"}.
+of \<open>c\<close> and \<open>d\<close>, and hence \<open>c mod d\<close> is \<open>0\<close> when exiting
+the loop. This implies that \<open>gcd c d = d\<close>, since \<open>c\<close> is divisible
+by \<open>d\<close>, so the conclusion follows using the assumption \<open>gcd c d = gcd m n\<close>.
This concludes the proofs of the open VCs, and hence the \SPARK{} verification
environment can be closed using the command \isa{\isacommand{spark\_end}}.
This command checks that all VCs have been proved and issues an error message
if there are remaining unproved VCs. Moreover, Isabelle checks that there is
no open \SPARK{} verification environment when the final \isa{\isacommand{end}}
command of a theory is encountered.
-*}
+\<close>
-section {* Optimizing the proof *}
+section \<open>Optimizing the proof\<close>
-text {*
+text \<open>
\begin{figure}
\lstinputlisting{Simple_Gcd.adb}
\input{Simple_Greatest_Common_Divisor}
@@ -227,8 +227,8 @@
numbers are non-negative by construction, the values computed by the algorithm
are trivially proved to be non-negative. Since we are working with non-negative
numbers, we can also just use \SPARK{}'s \textbf{mod} operator instead of
-\textbf{rem}, which spares us an application of theorems @{text zmod_zdiv_equality'}
-and @{text sdiv_pos_pos}. Finally, as noted by Barnes @{cite \<open>\S 11.5\<close> Barnes},
+\textbf{rem}, which spares us an application of theorems \<open>zmod_zdiv_equality'\<close>
+and \<open>sdiv_pos_pos\<close>. Finally, as noted by Barnes @{cite \<open>\S 11.5\<close> Barnes},
we can simplify matters by placing the \textbf{assert} statement between
\textbf{while} and \textbf{loop} rather than directly after the \textbf{loop}.
In the former case, the loop invariant has to be proved only once, whereas in
@@ -242,8 +242,8 @@
program for computing the greatest common divisor, together with its correctness
proof, is shown in \figref{fig:simple-gcd-proof}. Since the package specification
has not changed, we only show the body of the packages. The two VCs can now be
-proved by a single application of Isabelle's proof method @{text simp}.
-*}
+proved by a single application of Isabelle's proof method \<open>simp\<close>.
+\<close>
(*<*)
end
--- a/src/HOL/SPARK/Manual/Reference.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/SPARK/Manual/Reference.thy Thu May 26 17:51:22 2016 +0200
@@ -7,19 +7,19 @@
"_constrain" :: "logic => type => logic" ("_ :: _" [4, 0] 3)
(*>*)
-chapter {* HOL-\SPARK{} Reference *}
+chapter \<open>HOL-\SPARK{} Reference\<close>
-text {*
+text \<open>
\label{sec:spark-reference}
This section is intended as a quick reference for the HOL-\SPARK{} verification
environment. In \secref{sec:spark-commands}, we give a summary of the commands
provided by the HOL-\SPARK{}, while \secref{sec:spark-types} contains a description
of how particular types of \SPARK{} and FDL are modelled in Isabelle.
-*}
+\<close>
-section {* Commands *}
+section \<open>Commands\<close>
-text {*
+text \<open>
\label{sec:spark-commands}
This section describes the syntax and effect of each of the commands provided
by HOL-\SPARK{}.
@@ -87,43 +87,43 @@
otherwise the command issues an error message. As a side effect, the command
generates a proof review (\texttt{*.prv}) file to inform POGS of the proved
VCs.
-*}
+\<close>
-section {* Types *}
+section \<open>Types\<close>
-text {*
+text \<open>
\label{sec:spark-types}
The main types of FDL are integers, enumeration types, records, and arrays.
In the following sections, we describe how these types are modelled in
Isabelle.
-*}
+\<close>
-subsection {* Integers *}
+subsection \<open>Integers\<close>
-text {*
+text \<open>
The FDL type \texttt{integer} is modelled by the Isabelle type @{typ int}.
While the FDL \texttt{mod} operator behaves in the same way as its Isabelle
counterpart, this is not the case for the \texttt{div} operator. As has already
been mentioned in \secref{sec:proving-vcs}, the \texttt{div} operator of \SPARK{}
-always truncates towards zero, whereas the @{text div} operator of Isabelle
+always truncates towards zero, whereas the \<open>div\<close> operator of Isabelle
truncates towards minus infinity. Therefore, the FDL \texttt{div} operator is
-mapped to the @{text sdiv} operator in Isabelle. The characteristic theorems
-of @{text sdiv}, in particular those describing the relationship with the standard
-@{text div} operator, are shown in \figref{fig:sdiv-properties}
+mapped to the \<open>sdiv\<close> operator in Isabelle. The characteristic theorems
+of \<open>sdiv\<close>, in particular those describing the relationship with the standard
+\<open>div\<close> operator, are shown in \figref{fig:sdiv-properties}
\begin{figure}
\begin{center}
\small
\begin{tabular}{ll}
-@{text sdiv_def}: & @{thm sdiv_def} \\
-@{text sdiv_minus_dividend}: & @{thm sdiv_minus_dividend} \\
-@{text sdiv_minus_divisor}: & @{thm sdiv_minus_divisor} \\
-@{text sdiv_pos_pos}: & @{thm [mode=no_brackets] sdiv_pos_pos} \\
-@{text sdiv_pos_neg}: & @{thm [mode=no_brackets] sdiv_pos_neg} \\
-@{text sdiv_neg_pos}: & @{thm [mode=no_brackets] sdiv_neg_pos} \\
-@{text sdiv_neg_neg}: & @{thm [mode=no_brackets] sdiv_neg_neg} \\
+\<open>sdiv_def\<close>: & @{thm sdiv_def} \\
+\<open>sdiv_minus_dividend\<close>: & @{thm sdiv_minus_dividend} \\
+\<open>sdiv_minus_divisor\<close>: & @{thm sdiv_minus_divisor} \\
+\<open>sdiv_pos_pos\<close>: & @{thm [mode=no_brackets] sdiv_pos_pos} \\
+\<open>sdiv_pos_neg\<close>: & @{thm [mode=no_brackets] sdiv_pos_neg} \\
+\<open>sdiv_neg_pos\<close>: & @{thm [mode=no_brackets] sdiv_neg_pos} \\
+\<open>sdiv_neg_neg\<close>: & @{thm [mode=no_brackets] sdiv_neg_neg} \\
\end{tabular}
\end{center}
-\caption{Characteristic properties of @{text sdiv}}
+\caption{Characteristic properties of \<open>sdiv\<close>}
\label{fig:sdiv-properties}
\end{figure}
@@ -131,29 +131,29 @@
\begin{center}
\small
\begin{tabular}{ll}
-@{text AND_lower}: & @{thm [mode=no_brackets] AND_lower} \\
-@{text OR_lower}: & @{thm [mode=no_brackets] OR_lower} \\
-@{text XOR_lower}: & @{thm [mode=no_brackets] XOR_lower} \\
-@{text AND_upper1}: & @{thm [mode=no_brackets] AND_upper1} \\
-@{text AND_upper2}: & @{thm [mode=no_brackets] AND_upper2} \\
-@{text OR_upper}: & @{thm [mode=no_brackets] OR_upper} \\
-@{text XOR_upper}: & @{thm [mode=no_brackets] XOR_upper} \\
-@{text AND_mod}: & @{thm [mode=no_brackets] AND_mod}
+\<open>AND_lower\<close>: & @{thm [mode=no_brackets] AND_lower} \\
+\<open>OR_lower\<close>: & @{thm [mode=no_brackets] OR_lower} \\
+\<open>XOR_lower\<close>: & @{thm [mode=no_brackets] XOR_lower} \\
+\<open>AND_upper1\<close>: & @{thm [mode=no_brackets] AND_upper1} \\
+\<open>AND_upper2\<close>: & @{thm [mode=no_brackets] AND_upper2} \\
+\<open>OR_upper\<close>: & @{thm [mode=no_brackets] OR_upper} \\
+\<open>XOR_upper\<close>: & @{thm [mode=no_brackets] XOR_upper} \\
+\<open>AND_mod\<close>: & @{thm [mode=no_brackets] AND_mod}
\end{tabular}
\end{center}
\caption{Characteristic properties of bitwise operators}
\label{fig:bitwise}
\end{figure}
The bitwise logical operators of \SPARK{} and FDL are modelled by the operators
-@{text AND}, @{text OR} and @{text XOR} from Isabelle's @{text Word} library,
+\<open>AND\<close>, \<open>OR\<close> and \<open>XOR\<close> from Isabelle's \<open>Word\<close> library,
all of which have type @{typ "int \<Rightarrow> int \<Rightarrow> int"}. A list of properties of these
operators that are useful in proofs about \SPARK{} programs are shown in
\figref{fig:bitwise}
-*}
+\<close>
-subsection {* Enumeration types *}
+subsection \<open>Enumeration types\<close>
-text {*
+text \<open>
The FDL enumeration type
\begin{alltt}
type \(t\) = (\(e\sb{1}\), \(e\sb{2}\), \dots, \(e\sb{n}\));
@@ -165,7 +165,7 @@
\end{isabelle}
The HOL-\SPARK{} environment defines a type class @{class spark_enum} that captures
the characteristic properties of all enumeration types. It provides the following
-polymorphic functions and constants for all types @{text "'a"} of this type class:
+polymorphic functions and constants for all types \<open>'a\<close> of this type class:
\begin{flushleft}
@{term_type [mode=my_constrain] pos} \\
@{term_type [mode=my_constrain] val} \\
@@ -175,7 +175,7 @@
@{term_type [mode=my_constrain] last_el}
\end{flushleft}
In addition, @{class spark_enum} is a subclass of the @{class linorder} type class,
-which allows the comparison operators @{text "<"} and @{text "\<le>"} to be used on
+which allows the comparison operators \<open><\<close> and \<open>\<le>\<close> to be used on
enumeration types. The polymorphic operations shown above enjoy a number of
generic properties that hold for all enumeration types. These properties are
listed in \figref{fig:enum-generic-properties}.
@@ -186,23 +186,23 @@
\begin{center}
\small
\begin{tabular}{ll}
-@{text range_pos}: & @{thm range_pos} \\
-@{text less_pos}: & @{thm less_pos} \\
-@{text less_eq_pos}: & @{thm less_eq_pos} \\
-@{text val_def}: & @{thm val_def} \\
-@{text succ_def}: & @{thm succ_def} \\
-@{text pred_def}: & @{thm pred_def} \\
-@{text first_el_def}: & @{thm first_el_def} \\
-@{text last_el_def}: & @{thm last_el_def} \\
-@{text inj_pos}: & @{thm inj_pos} \\
-@{text val_pos}: & @{thm val_pos} \\
-@{text pos_val}: & @{thm pos_val} \\
-@{text first_el_smallest}: & @{thm first_el_smallest} \\
-@{text last_el_greatest}: & @{thm last_el_greatest} \\
-@{text pos_succ}: & @{thm pos_succ} \\
-@{text pos_pred}: & @{thm pos_pred} \\
-@{text succ_val}: & @{thm succ_val} \\
-@{text pred_val}: & @{thm pred_val}
+\<open>range_pos\<close>: & @{thm range_pos} \\
+\<open>less_pos\<close>: & @{thm less_pos} \\
+\<open>less_eq_pos\<close>: & @{thm less_eq_pos} \\
+\<open>val_def\<close>: & @{thm val_def} \\
+\<open>succ_def\<close>: & @{thm succ_def} \\
+\<open>pred_def\<close>: & @{thm pred_def} \\
+\<open>first_el_def\<close>: & @{thm first_el_def} \\
+\<open>last_el_def\<close>: & @{thm last_el_def} \\
+\<open>inj_pos\<close>: & @{thm inj_pos} \\
+\<open>val_pos\<close>: & @{thm val_pos} \\
+\<open>pos_val\<close>: & @{thm pos_val} \\
+\<open>first_el_smallest\<close>: & @{thm first_el_smallest} \\
+\<open>last_el_greatest\<close>: & @{thm last_el_greatest} \\
+\<open>pos_succ\<close>: & @{thm pos_succ} \\
+\<open>pos_pred\<close>: & @{thm pos_pred} \\
+\<open>succ_val\<close>: & @{thm succ_val} \\
+\<open>pred_val\<close>: & @{thm pred_val}
\end{tabular}
\end{center}
\caption{Generic properties of functions on enumeration types}
@@ -226,11 +226,11 @@
\caption{Type-specific properties of functions on enumeration types}
\label{fig:enum-specific-properties}
\end{figure}
-*}
+\<close>
-subsection {* Records *}
+subsection \<open>Records\<close>
-text {*
+text \<open>
The FDL record type
\begin{alltt}
type \(t\) = record
@@ -252,11 +252,11 @@
a field $f_i$ of a record $r$ is selected using the notation $f_i~r$, and the
fields $f$ and $f'$ of a record $r$ can be updated using the notation
\mbox{\isa{$r$\ \isasymlparr$f$\ :=\ $v$,\ $f'$\ :=\ $v'$\isasymrparr}}.
-*}
+\<close>
-subsection {* Arrays *}
+subsection \<open>Arrays\<close>
-text {*
+text \<open>
The FDL array type
\begin{alltt}
type \(t\) = array [\(t\sb{1}\), \(\ldots\), \(t\sb{n}\)] of \(u\);
@@ -270,11 +270,11 @@
Thus, we can write expressions like
@{term [display] "(A::int\<Rightarrow>int) ({0..9} [:=] 42, 15 := 99, {20..29} [:=] 0)"}
that would be cumbersome to write using single updates.
-*}
+\<close>
-section {* User-defined proof functions and types *}
+section \<open>User-defined proof functions and types\<close>
-text {*
+text \<open>
To illustrate the interplay between the commands for introducing user-defined proof
functions and types mentioned in \secref{sec:spark-commands}, we now discuss a larger
example involving the definition of proof functions on complex types. Assume we would
@@ -301,8 +301,8 @@
functions involving the enumeration and record types introduced above, and link
them to the corresponding \SPARK{} proof functions. It is important that the
\isa{\isacommand{definition}} commands are preceeded by the \isa{\isacommand{spark\_types}}
-command, since the definition of @{text initialized3} uses the @{text val}
-function for enumeration types that is only available once that @{text day}
+command, since the definition of \<open>initialized3\<close> uses the \<open>val\<close>
+function for enumeration types that is only available once that \<open>day\<close>
has been declared as a \SPARK{} type.
\begin{figure}
\lstinputlisting{complex_types.ads}
@@ -320,7 +320,7 @@
\caption{Theory defining proof functions for complex types}
\label{fig:complex-types-thy}
\end{figure}
-*}
+\<close>
(*<*)
end
--- a/src/HOL/SPARK/Manual/Simple_Greatest_Common_Divisor.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/SPARK/Manual/Simple_Greatest_Common_Divisor.thy Thu May 26 17:51:22 2016 +0200
@@ -13,11 +13,11 @@
spark_open "simple_greatest_common_divisor/g_c_d"
spark_vc procedure_g_c_d_4
- using `0 < d` `gcd c d = gcd m n`
+ using \<open>0 < d\<close> \<open>gcd c d = gcd m n\<close>
by (simp add: gcd_non_0_int)
spark_vc procedure_g_c_d_9
- using `0 \<le> c` `gcd c 0 = gcd m n`
+ using \<open>0 \<le> c\<close> \<open>gcd c 0 = gcd m n\<close>
by simp
spark_end
--- a/src/HOL/SPARK/Manual/VC_Principles.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/SPARK/Manual/VC_Principles.thy Thu May 26 17:51:22 2016 +0200
@@ -4,9 +4,9 @@
begin
(*>*)
-chapter {* Principles of VC generation *}
+chapter \<open>Principles of VC generation\<close>
-text {*
+text \<open>
\label{sec:vc-principles}
In this section, we will discuss some aspects of VC generation that are
useful for understanding and optimizing the VCs produced by the \SPARK{}
@@ -145,7 +145,7 @@
The VC for the path from assertion 2 to assertion 1 is trivial, and so is the VC for the
path from assertion 2 to the postcondition, expressing that the loop invariant implies
the postcondition when the loop has terminated.
-*}
+\<close>
(*<*)
end
--- a/src/HOL/SPARK/SPARK.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/SPARK/SPARK.thy Thu May 26 17:51:22 2016 +0200
@@ -9,7 +9,7 @@
imports SPARK_Setup
begin
-text {* Bitwise logical operators *}
+text \<open>Bitwise logical operators\<close>
spark_proof_functions
bit__and (integer, integer) : integer = "op AND"
@@ -54,7 +54,7 @@
bit_not_spark_eq [where 'a=64, simplified]
-text {* Minimum and maximum *}
+text \<open>Minimum and maximum\<close>
spark_proof_functions
integer__min = "min :: int \<Rightarrow> int \<Rightarrow> int"
--- a/src/HOL/SPARK/SPARK_Setup.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/SPARK/SPARK_Setup.thy Thu May 26 17:51:22 2016 +0200
@@ -17,9 +17,9 @@
ML_file "Tools/fdl_lexer.ML"
ML_file "Tools/fdl_parser.ML"
-text {*
+text \<open>
SPARK version of div, see section 4.4.1.1 of SPARK Proof Manual
-*}
+\<close>
definition sdiv :: "int \<Rightarrow> int \<Rightarrow> int" (infixl "sdiv" 70) where
"a sdiv b = sgn a * sgn b * (\<bar>a\<bar> div \<bar>b\<bar>)"
@@ -30,9 +30,9 @@
lemma sdiv_minus_divisor: "a sdiv - b = - (a sdiv b)"
by (simp add: sdiv_def sgn_if)
-text {*
+text \<open>
Correspondence between HOL's and SPARK's version of div
-*}
+\<close>
lemma sdiv_pos_pos: "0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> a sdiv b = a div b"
by (simp add: sdiv_def sgn_if)
@@ -47,9 +47,9 @@
by (simp add: sdiv_def sgn_if)
-text {*
+text \<open>
Updating a function at a set of points. Useful for building arrays.
-*}
+\<close>
definition fun_upds :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> 'b" where
"fun_upds f xs y z = (if z \<in> xs then y else f z)"
@@ -70,7 +70,7 @@
by (simp add: fun_eq_iff)
-text {* Enumeration types *}
+text \<open>Enumeration types\<close>
class spark_enum = ord + finite +
fixes pos :: "'a \<Rightarrow> int"
@@ -178,7 +178,7 @@
by auto
-text {* Load the package *}
+text \<open>Load the package\<close>
ML_file "Tools/spark_vcs.ML"
ML_file "Tools/spark_commands.ML"
--- a/src/HOL/Statespace/DistinctTreeProver.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Statespace/DistinctTreeProver.thy Thu May 26 17:51:22 2016 +0200
@@ -2,33 +2,33 @@
Author: Norbert Schirmer, TU Muenchen
*)
-section {* Distinctness of Names in a Binary Tree \label{sec:DistinctTreeProver}*}
+section \<open>Distinctness of Names in a Binary Tree \label{sec:DistinctTreeProver}\<close>
theory DistinctTreeProver
imports Main
begin
-text {* A state space manages a set of (abstract) names and assumes
+text \<open>A state space manages a set of (abstract) names and assumes
that the names are distinct. The names are stored as parameters of a
locale and distinctness as an assumption. The most common request is
to proof distinctness of two given names. We maintain the names in a
balanced binary tree and formulate a predicate that all nodes in the
tree have distinct names. This setup leads to logarithmic certificates.
-*}
+\<close>
-subsection {* The Binary Tree *}
+subsection \<open>The Binary Tree\<close>
datatype 'a tree = Node "'a tree" 'a bool "'a tree" | Tip
-text {* The boolean flag in the node marks the content of the node as
+text \<open>The boolean flag in the node marks the content of the node as
deleted, without having to build a new tree. We prefer the boolean
flag to an option type, so that the ML-layer can still use the node
content to facilitate binary search in the tree. The ML code keeps the
nodes sorted using the term order. We do not have to push ordering to
-the HOL level. *}
+the HOL level.\<close>
-subsection {* Distinctness of Nodes *}
+subsection \<open>Distinctness of Nodes\<close>
primrec set_of :: "'a tree \<Rightarrow> 'a set"
@@ -44,11 +44,11 @@
set_of l \<inter> set_of r = {} \<and>
all_distinct l \<and> all_distinct r)"
-text {* Given a binary tree @{term "t"} for which
+text \<open>Given a binary tree @{term "t"} for which
@{const all_distinct} holds, given two different nodes contained in the tree,
we want to write a ML function that generates a logarithmic
certificate that the content of the nodes is distinct. We use the
-following lemmas to achieve this. *}
+following lemmas to achieve this.\<close>
lemma all_distinct_left: "all_distinct (Node l x b r) \<Longrightarrow> all_distinct l"
by simp
@@ -81,9 +81,9 @@
lemma neq_to_eq_False: "x\<noteq>y \<Longrightarrow> (x=y)\<equiv>False"
by simp
-subsection {* Containment of Trees *}
+subsection \<open>Containment of Trees\<close>
-text {* When deriving a state space from other ones, we create a new
+text \<open>When deriving a state space from other ones, we create a new
name tree which contains all the names of the parent state spaces and
assume the predicate @{const all_distinct}. We then prove that the new
locale interprets all parent locales. Hence we have to show that the
@@ -95,7 +95,7 @@
implies @{const all_distinct} of the parent tree. The resulting
certificate is of the order @{term "n * log(m)"} where @{term "n"} is
the size of the (smaller) parent tree and @{term "m"} the size of the
-(bigger) new tree. *}
+(bigger) new tree.\<close>
primrec delete :: "'a \<Rightarrow> 'a tree \<Rightarrow> 'a tree option"
@@ -627,8 +627,8 @@
lemma subtract_Tip: "subtract Tip t = Some t"
by simp
-text {* Now we have all the theorems in place that are needed for the
-certificate generating ML functions. *}
+text \<open>Now we have all the theorems in place that are needed for the
+certificate generating ML functions.\<close>
ML_file "distinct_tree_prover.ML"
--- a/src/HOL/Statespace/StateFun.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Statespace/StateFun.thy Thu May 26 17:51:22 2016 +0200
@@ -2,13 +2,13 @@
Author: Norbert Schirmer, TU Muenchen
*)
-section {* State Space Representation as Function \label{sec:StateFun}*}
+section \<open>State Space Representation as Function \label{sec:StateFun}\<close>
theory StateFun imports DistinctTreeProver
begin
-text {* The state space is represented as a function from names to
+text \<open>The state space is represented as a function from names to
values. We neither fix the type of names nor the type of values. We
define lookup and update functions and provide simprocs that simplify
expressions containing these, similar to HOL-records.
@@ -20,7 +20,7 @@
The update is actually generalized to a map function. The map supplies
better compositionality, especially if you think of nested state
-spaces. *}
+spaces.\<close>
definition K_statefun :: "'a \<Rightarrow> 'b \<Rightarrow> 'a" where "K_statefun c x \<equiv> c"
--- a/src/HOL/Statespace/StateSpaceEx.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Statespace/StateSpaceEx.thy Thu May 26 17:51:22 2016 +0200
@@ -2,7 +2,7 @@
Author: Norbert Schirmer, TU Muenchen
*)
-section {* Examples \label{sec:Examples} *}
+section \<open>Examples \label{sec:Examples}\<close>
theory StateSpaceEx
imports StateSpaceLocale StateSpaceSyntax
begin
@@ -12,14 +12,14 @@
"_statespace_updates" :: "('a \<Rightarrow> 'b) \<Rightarrow> updbinds \<Rightarrow> ('a \<Rightarrow> 'b)" ("_\<langle>_\<rangle>" [900,0] 900)
(*>*)
-text {* Did you ever dream about records with multiple inheritance?
+text \<open>Did you ever dream about records with multiple inheritance?
Then you should definitely have a look at statespaces. They may be
-what you are dreaming of. Or at least almost \dots *}
+what you are dreaming of. Or at least almost \dots\<close>
-text {* Isabelle allows to add new top-level commands to the
+text \<open>Isabelle allows to add new top-level commands to the
system. Building on the locale infrastructure, we provide a command
-\<^theory_text>\<open>statespace\<close> like this:*}
+\<^theory_text>\<open>statespace\<close> like this:\<close>
statespace vars =
n::nat
@@ -29,12 +29,12 @@
print_locale vars_valuetypes
print_locale vars
-text {* \noindent This resembles a \<^theory_text>\<open>record\<close> definition,
+text \<open>\noindent This resembles a \<^theory_text>\<open>record\<close> definition,
but introduces sophisticated locale
infrastructure instead of HOL type schemes. The resulting context
postulates two distinct names @{term "n"} and @{term "b"} and
projection~/ injection functions that convert from abstract values to
-@{typ "nat"} and @{text "bool"}. The logical content of the locale is: *}
+@{typ "nat"} and \<open>bool\<close>. The logical content of the locale is:\<close>
locale vars' =
fixes n::'name and b::'name
@@ -46,8 +46,8 @@
fixes project_bool::"'value \<Rightarrow> bool" and inject_bool::"bool \<Rightarrow> 'value"
assumes "\<And>b. project_bool (inject_bool b) = b"
-text {* \noindent The HOL predicate @{const "distinct"} describes
-distinctness of all names in the context. Locale @{text "vars'"}
+text \<open>\noindent The HOL predicate @{const "distinct"} describes
+distinctness of all names in the context. Locale \<open>vars'\<close>
defines the raw logical content that is defined in the state space
locale. We also maintain non-logical context information to support
the user:
@@ -71,44 +71,43 @@
extended there are multiple distinctness theorems in the context. Our
declarations take care that the link always points to the strongest
distinctness assumption. With these declarations in place, a lookup
-can be written as @{text "s\<cdot>n"}, which is translated to @{text
-"project_nat (s n)"}, and an update as @{text "s\<langle>n := 2\<rangle>"}, which is
-translated to @{text "s(n := inject_nat 2)"}. We can now establish the
-following lemma: *}
+can be written as \<open>s\<cdot>n\<close>, which is translated to \<open>project_nat (s n)\<close>, and an update as \<open>s\<langle>n := 2\<rangle>\<close>, which is
+translated to \<open>s(n := inject_nat 2)\<close>. We can now establish the
+following lemma:\<close>
lemma (in vars) foo: "s<n := 2>\<cdot>b = s\<cdot>b" by simp
-text {* \noindent Here the simplifier was able to refer to
+text \<open>\noindent Here the simplifier was able to refer to
distinctness of @{term "b"} and @{term "n"} to solve the equation.
-The resulting lemma is also recorded in locale @{text "vars"} for
+The resulting lemma is also recorded in locale \<open>vars\<close> for
later use and is automatically propagated to all its interpretations.
-Here is another example: *}
+Here is another example:\<close>
statespace 'a varsX = NB: vars [n=N, b=B] + vars + x::'a
-text {* \noindent The state space @{text "varsX"} imports two copies
-of the state space @{text "vars"}, where one has the variables renamed
+text \<open>\noindent The state space \<open>varsX\<close> imports two copies
+of the state space \<open>vars\<close>, where one has the variables renamed
to upper-case letters, and adds another variable @{term "x"} of type
@{typ "'a"}. This type is fixed inside the state space but may get
instantiated later on, analogous to type parameters of an ML-functor.
-The distinctness assumption is now @{text "distinct [N, B, n, b, x]"},
+The distinctness assumption is now \<open>distinct [N, B, n, b, x]\<close>,
from this we can derive both @{term "distinct [N,B]"} and @{term
"distinct [n,b]"}, the distinction assumptions for the two versions of
-locale @{text "vars"} above. Moreover we have all necessary
+locale \<open>vars\<close> above. Moreover we have all necessary
projection and injection assumptions available. These assumptions
together allow us to establish state space @{term "varsX"} as an
interpretation of both instances of locale @{term "vars"}. Hence we
-inherit both variants of theorem @{text "foo"}: @{text "s\<langle>N := 2\<rangle>\<cdot>B =
-s\<cdot>B"} as well as @{text "s\<langle>n := 2\<rangle>\<cdot>b = s\<cdot>b"}. These are immediate
+inherit both variants of theorem \<open>foo\<close>: \<open>s\<langle>N := 2\<rangle>\<cdot>B =
+s\<cdot>B\<close> as well as \<open>s\<langle>n := 2\<rangle>\<cdot>b = s\<cdot>b\<close>. These are immediate
consequences of the locale interpretation action.
The declarations for syntax and the distinctness theorems also observe
the morphisms generated by the locale package due to the renaming
-@{term "n = N"}: *}
+@{term "n = N"}:\<close>
lemma (in varsX) foo: "s\<langle>N := 2\<rangle>\<cdot>x = s\<cdot>x" by simp
-text {* To assure scalability towards many distinct names, the
+text \<open>To assure scalability towards many distinct names, the
distinctness predicate is refined to operate on balanced trees. Thus
we get logarithmic certificates for the distinctness of two names by
the distinctness of the paths in the tree. Asked for the distinctness
@@ -118,9 +117,9 @@
spaces requires to prove that the combined distinctness assumption
implies the distinctness assumptions of the components. Such a proof
is of the order $m \cdot \log n$, where $n$ and $m$ are the number of
-nodes in the larger and smaller tree, respectively.*}
+nodes in the larger and smaller tree, respectively.\<close>
-text {* We continue with more examples. *}
+text \<open>We continue with more examples.\<close>
statespace 'a foo =
f::"nat\<Rightarrow>nat"
@@ -154,9 +153,9 @@
shows "(s\<langle>b:=True\<rangle>)\<cdot>c = s\<cdot>c"
by simp
-text {* You can define a derived state space by inheriting existing state spaces, renaming
+text \<open>You can define a derived state space by inheriting existing state spaces, renaming
of components if you like, and by declaring new components.
-*}
+\<close>
statespace ('a,'b) loo = 'a foo + bar [b=B,c=C] +
X::'b
@@ -165,13 +164,13 @@
shows "s\<langle>a:=i\<rangle>\<cdot>B = s\<cdot>B"
proof -
thm foo1
- txt {* The Lemma @{thm [source] foo1} from the parent state space
- is also available here: \begin{center}@{thm foo1}\end{center} *}
+ txt \<open>The Lemma @{thm [source] foo1} from the parent state space
+ is also available here: \begin{center}@{thm foo1}\end{center}\<close>
have "s<a:=i>\<cdot>a = i"
by (rule foo1)
thm bar1
- txt {* Note the renaming of the parameters in Lemma @{thm [source] bar1}:
- \begin{center}@{thm bar1}\end{center} *}
+ txt \<open>Note the renaming of the parameters in Lemma @{thm [source] bar1}:
+ \begin{center}@{thm bar1}\end{center}\<close>
have "s<B:=True>\<cdot>C = s\<cdot>C"
by (rule bar1)
show ?thesis
@@ -203,9 +202,9 @@
*)
(* ++++++++++++++++++++++++++++++++++++++++++++++++++++++++++++++++++++++++ *)
-text {* There are known problems with syntax-declarations. They currently
+text \<open>There are known problems with syntax-declarations. They currently
only work, when the context is already built. Hopefully this will be
-implemented correctly in future Isabelle versions. *}
+implemented correctly in future Isabelle versions.\<close>
(*
lemma
@@ -222,23 +221,23 @@
shows "s<a := i>\<cdot>a = i"
*)
-text {* It would be nice to have nested state spaces. This is
+text \<open>It would be nice to have nested state spaces. This is
logically no problem. From the locale-implementation side this may be
something like an 'includes' into a locale. When there is a more
elaborate locale infrastructure in place this may be an easy exercise.
-*}
+\<close>
-subsection {* Benchmarks *}
+subsection \<open>Benchmarks\<close>
-text {* Here are some bigger examples for benchmarking. *}
+text \<open>Here are some bigger examples for benchmarking.\<close>
-ML {*
+ML \<open>
fun make_benchmark n =
writeln (Active.sendback_markup []
("statespace benchmark" ^ string_of_int n ^ " =\n" ^
(cat_lines (map (fn i => "A" ^ string_of_int i ^ "::nat") (1 upto n)))));
-*}
+\<close>
text "0.2s"
statespace benchmark100 = A1::nat A2::nat A3::nat A4::nat A5::nat
--- a/src/HOL/Statespace/StateSpaceLocale.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Statespace/StateSpaceLocale.thy Thu May 26 17:51:22 2016 +0200
@@ -2,7 +2,7 @@
Author: Norbert Schirmer, TU Muenchen
*)
-section {* Setup for State Space Locales \label{sec:StateSpaceLocale}*}
+section \<open>Setup for State Space Locales \label{sec:StateSpaceLocale}\<close>
theory StateSpaceLocale imports StateFun
keywords "statespace" :: thy_decl
@@ -11,9 +11,9 @@
ML_file "state_space.ML"
ML_file "state_fun.ML"
-text {* For every type that is to be stored in a state space, an
+text \<open>For every type that is to be stored in a state space, an
instance of this locale is imported in order convert the abstract and
-concrete values.*}
+concrete values.\<close>
locale project_inject =
--- a/src/HOL/Statespace/StateSpaceSyntax.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/Statespace/StateSpaceSyntax.thy Thu May 26 17:51:22 2016 +0200
@@ -2,13 +2,13 @@
Author: Norbert Schirmer, TU Muenchen
*)
-section {* Syntax for State Space Lookup and Update \label{sec:StateSpaceSyntax}*}
+section \<open>Syntax for State Space Lookup and Update \label{sec:StateSpaceSyntax}\<close>
theory StateSpaceSyntax
imports StateSpaceLocale
begin
-text {* The state space syntax is kept in an extra theory so that you
-can choose if you want to use it or not. *}
+text \<open>The state space syntax is kept in an extra theory so that you
+can choose if you want to use it or not.\<close>
syntax
"_statespace_lookup" :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'c" ("_\<cdot>_" [60, 60] 60)
@@ -22,16 +22,16 @@
parse_translation
-{*
+\<open>
[(@{syntax_const "_statespace_lookup"}, StateSpace.lookup_tr),
(@{syntax_const "_statespace_update"}, StateSpace.update_tr)]
-*}
+\<close>
print_translation
-{*
+\<open>
[(@{const_syntax lookup}, StateSpace.lookup_tr'),
(@{const_syntax update}, StateSpace.update_tr')]
-*}
+\<close>
end
--- a/src/HOL/TPTP/ATP_Problem_Import.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/TPTP/ATP_Problem_Import.thy Thu May 26 17:51:22 2016 +0200
@@ -2,7 +2,7 @@
Author: Jasmin Blanchette, TU Muenchen
*)
-section {* ATP Problem Importer *}
+section \<open>ATP Problem Importer\<close>
theory ATP_Problem_Import
imports Complex_Main TPTP_Interpret "~~/src/HOL/Library/Refute"
@@ -10,7 +10,7 @@
ML_file "sledgehammer_tactics.ML"
-ML {* Proofterm.proofs := 0 *}
+ML \<open>Proofterm.proofs := 0\<close>
declare [[show_consts]] (* for Refute *)
declare [[smt_oracle]]
--- a/src/HOL/TPTP/ATP_Theory_Export.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/TPTP/ATP_Theory_Export.thy Thu May 26 17:51:22 2016 +0200
@@ -2,7 +2,7 @@
Author: Jasmin Blanchette, TU Muenchen
*)
-section {* ATP Theory Exporter *}
+section \<open>ATP Theory Exporter\<close>
theory ATP_Theory_Export
imports Complex_Main
@@ -10,61 +10,61 @@
ML_file "atp_theory_export.ML"
-ML {*
+ML \<open>
open ATP_Problem
open ATP_Theory_Export
-*}
+\<close>
-ML {*
+ML \<open>
val do_it = false (* switch to "true" to generate the files *)
val ctxt = @{context}
val thy = @{theory Complex_Main}
val infer_policy = (* Unchecked_Inferences *) No_Inferences
-*}
+\<close>
-ML {*
+ML \<open>
if do_it then
"/tmp/isa_filter"
|> generate_casc_lbt_isa_files_for_theory ctxt thy (THF (Polymorphic, THF_Without_Choice))
infer_policy "poly_native_higher"
else
()
-*}
+\<close>
-ML {*
+ML \<open>
if do_it then
"/tmp/axs_tc_native.dfg"
|> generate_atp_inference_file_for_theory ctxt thy (DFG Polymorphic)
infer_policy "tc_native"
else
()
-*}
+\<close>
-ML {*
+ML \<open>
if do_it then
"/tmp/infs_poly_guards_query_query.tptp"
|> generate_atp_inference_file_for_theory ctxt thy FOF infer_policy
"poly_guards??"
else
()
-*}
+\<close>
-ML {*
+ML \<open>
if do_it then
"/tmp/infs_poly_tags_query_query.tptp"
|> generate_atp_inference_file_for_theory ctxt thy FOF infer_policy
"poly_tags??"
else
()
-*}
+\<close>
-ML {*
+ML \<open>
if do_it then
"/tmp/casc_ltb_isa"
|> generate_casc_lbt_isa_files_for_theory ctxt thy FOF infer_policy
"poly_tags??"
else
()
-*}
+\<close>
end
--- a/src/HOL/TPTP/MaSh_Eval.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/TPTP/MaSh_Eval.thy Thu May 26 17:51:22 2016 +0200
@@ -2,7 +2,7 @@
Author: Jasmin Blanchette, TU Muenchen
*)
-section {* MaSh Evaluation Driver *}
+section \<open>MaSh Evaluation Driver\<close>
theory MaSh_Eval
imports MaSh_Export_Base
@@ -14,27 +14,27 @@
[provers = e, max_facts = 64, strict, dont_slice, type_enc = poly_guards??,
lam_trans = combs, timeout = 30, dont_preplay, minimize]
-ML {*
+ML \<open>
Multithreading.max_threads ()
-*}
+\<close>
-ML {*
+ML \<open>
open MaSh_Eval
-*}
+\<close>
-ML {*
+ML \<open>
val params = Sledgehammer_Commands.default_params @{theory} []
val prob_dir = prefix ^ "mash_problems"
-*}
+\<close>
-ML {*
+ML \<open>
if do_it then
Isabelle_System.mkdir (Path.explode prob_dir)
else
()
-*}
+\<close>
-ML {*
+ML \<open>
if do_it then
evaluate_mash_suggestions @{context} params range (SOME prob_dir)
[prefix ^ "mepo_suggestions",
@@ -45,6 +45,6 @@
(prefix ^ "mash_eval")
else
()
-*}
+\<close>
end
--- a/src/HOL/TPTP/MaSh_Export.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/TPTP/MaSh_Export.thy Thu May 26 17:51:22 2016 +0200
@@ -2,126 +2,126 @@
Author: Jasmin Blanchette, TU Muenchen
*)
-section {* MaSh Exporter *}
+section \<open>MaSh Exporter\<close>
theory MaSh_Export
imports MaSh_Export_Base
begin
-ML {*
+ML \<open>
if do_it then
Isabelle_System.mkdir (Path.explode prefix)
else
()
-*}
+\<close>
-ML {*
+ML \<open>
if do_it then
generate_mash_suggestions "nb" @{context} params (range, step) thys max_suggestions
(prefix ^ "mash_nb_suggestions")
else
()
-*}
+\<close>
-ML {*
+ML \<open>
if do_it then
generate_mash_suggestions "knn" @{context} params (range, step) thys max_suggestions
(prefix ^ "mash_knn_suggestions")
else
()
-*}
+\<close>
-ML {*
+ML \<open>
if do_it then
generate_mepo_suggestions @{context} params (range, step) thys max_suggestions
(prefix ^ "mepo_suggestions")
else
()
-*}
+\<close>
-ML {*
+ML \<open>
if do_it then
generate_mesh_suggestions max_suggestions (prefix ^ "mash_nb_suggestions")
(prefix ^ "mepo_suggestions") (prefix ^ "mesh_nb_suggestions")
else
()
-*}
+\<close>
-ML {*
+ML \<open>
if do_it then
generate_mesh_suggestions max_suggestions (prefix ^ "mash_knn_suggestions")
(prefix ^ "mepo_suggestions") (prefix ^ "mesh_knn_suggestions")
else
()
-*}
+\<close>
-ML {*
+ML \<open>
if do_it then
generate_prover_dependencies @{context} params range thys
(prefix ^ "mash_nb_prover_dependencies")
else
()
-*}
+\<close>
-ML {*
+ML \<open>
if do_it then
generate_prover_dependencies @{context} params range thys
(prefix ^ "mash_knn_prover_dependencies")
else
()
-*}
+\<close>
-ML {*
+ML \<open>
if do_it then
generate_mesh_suggestions max_suggestions (prefix ^ "mash_nb_prover_suggestions")
(prefix ^ "mepo_suggestions") (prefix ^ "mesh_nb_prover_suggestions")
else
()
-*}
+\<close>
-ML {*
+ML \<open>
if do_it then
generate_mesh_suggestions max_suggestions (prefix ^ "mash_knn_prover_suggestions")
(prefix ^ "mepo_suggestions") (prefix ^ "mesh_knn_prover_suggestions")
else
()
-*}
+\<close>
-ML {*
+ML \<open>
if do_it then
generate_accessibility @{context} thys (prefix ^ "mash_accessibility")
else
()
-*}
+\<close>
-ML {*
+ML \<open>
if do_it then
generate_features @{context} thys (prefix ^ "mash_features")
else
()
-*}
+\<close>
-ML {*
+ML \<open>
if do_it then
generate_isar_dependencies @{context} range thys (prefix ^ "mash_dependencies")
else
()
-*}
+\<close>
-ML {*
+ML \<open>
if do_it then
generate_isar_commands @{context} prover (range, step) thys max_suggestions
(prefix ^ "mash_commands")
else
()
-*}
+\<close>
-ML {*
+ML \<open>
if do_it then
generate_prover_commands @{context} params (range, step) thys max_suggestions
(prefix ^ "mash_prover_commands")
else
()
-*}
+\<close>
end
--- a/src/HOL/TPTP/MaSh_Export_Base.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/TPTP/MaSh_Export_Base.thy Thu May 26 17:51:22 2016 +0200
@@ -2,7 +2,7 @@
Author: Jasmin Blanchette, TU Muenchen
*)
-section {* MaSh Exporter Base *}
+section \<open>MaSh Exporter Base\<close>
theory MaSh_Export_Base
imports Main
@@ -19,15 +19,15 @@
hide_fact (open) HOL.ext
-ML {*
+ML \<open>
Multithreading.max_threads ()
-*}
+\<close>
-ML {*
+ML \<open>
open MaSh_Export
-*}
+\<close>
-ML {*
+ML \<open>
val do_it = false (* switch to "true" to generate the files *)
val thys = [@{theory List}]
val params as {provers, ...} = Sledgehammer_Commands.default_params @{theory} []
@@ -37,6 +37,6 @@
val max_suggestions = 1024
val dir = "List"
val prefix = "/tmp/" ^ dir ^ "/"
-*}
+\<close>
end
--- a/src/HOL/TPTP/TPTP_Interpret_Test.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/TPTP/TPTP_Interpret_Test.thy Thu May 26 17:51:22 2016 +0200
@@ -12,7 +12,7 @@
section "Interpreter tests"
text "Interpret a problem."
-ML {*
+ML \<open>
val (time, ((type_map, const_map, fmlas), thy)) =
Timing.timing
(TPTP_Interpret.interpret_file
@@ -22,17 +22,17 @@
[]
[])
@{theory}
-*}
+\<close>
text "... and display nicely."
-ML {*
+ML \<open>
List.app (Pretty.writeln o (Syntax.pretty_term @{context}) o #3) fmlas;
-*}
+\<close>
subsection "Multiple tests"
-ML {*
+ML \<open>
(*default timeout is 1 min*)
fun interpret timeout file thy =
Timeout.apply (Time.fromSeconds (if timeout = 0 then 60 else timeout))
@@ -56,9 +56,9 @@
List.app
(interpretation_test timeout ctxt)
(map situate probs)
-*}
+\<close>
-ML {*
+ML \<open>
val some_probs =
["LCL/LCL825-1.p",
"ALG/ALG001^5.p",
@@ -111,19 +111,19 @@
"SYO/SYO561_2.p",
"SYO/SYO562_1.p",
"SYN/SYN000_2.p"]
-*}
+\<close>
-ML {*
+ML \<open>
interpretation_tests (get_timeout @{context}) @{context}
(some_probs @ take_too_long @ timeouts @ more_probs)
-*}
+\<close>
-ML {*
+ML \<open>
parse_timed (situate "NUM/NUM923^4.p");
interpret_timed 0 (situate "NUM/NUM923^4.p") @{theory}
-*}
+\<close>
-ML {*
+ML \<open>
fun interp_gain timeout thy file =
let
val t1 = (parse_timed file |> fst)
@@ -144,18 +144,18 @@
diff_elapsed / elapsed,
elapsed)
end
-*}
+\<close>
subsection "Test against whole TPTP"
text "Run interpretation over all problems. This works only for logics
for which interpretation is defined (in TPTP_Parser/tptp_interpret.ML)."
-ML {*
+ML \<open>
if test_all @{context} then
(report @{context} "Interpreting all problems";
S timed_test (interpretation_test (get_timeout @{context})) @{context} (TPTP_Syntax.get_file_list tptp_probs_dir))
else ()
-*}
+\<close>
end
--- a/src/HOL/TPTP/TPTP_Parser.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/TPTP/TPTP_Parser.thy Thu May 26 17:51:22 2016 +0200
@@ -16,13 +16,13 @@
ML_file "TPTP_Parser/tptp_problem_name.ML"
ML_file "TPTP_Parser/tptp_proof.ML"
-text {*The TPTP parser was generated using ML-Yacc, and needs the
+text \<open>The TPTP parser was generated using ML-Yacc, and needs the
ML-Yacc library to operate. This library is included with the parser,
and we include the next section in accordance with ML-Yacc's terms of
-use.*}
+use.\<close>
section "ML-YACC COPYRIGHT NOTICE, LICENSE AND DISCLAIMER."
-text {*
+text \<open>
Copyright 1989, 1990 by David R. Tarditi Jr. and Andrew W. Appel
Permission to use, copy, modify, and distribute this software and its
@@ -42,6 +42,6 @@
use, data or profits, whether in an action of contract, negligence or
other tortious action, arising out of or in connection with the use or
performance of this software.
-*}
+\<close>
end
--- a/src/HOL/TPTP/TPTP_Parser_Example.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/TPTP/TPTP_Parser_Example.thy Thu May 26 17:51:22 2016 +0200
@@ -12,24 +12,24 @@
import_tptp "$TPTP/Problems/LCL/LCL414+1.p"
-ML {*
+ML \<open>
val an_fmlas =
TPTP_Interpret.get_manifests @{theory}
|> hd (*FIXME use named lookup*)
|> #2 (*get problem contents*)
|> #3 (*get formulas*)
-*}
+\<close>
(*Display nicely.*)
-ML {*
+ML \<open>
List.app (fn (n, role, fmla, _) =>
Pretty.writeln
(Pretty.block [Pretty.str ("\"" ^ n ^ "\"" ^ "(" ^
TPTP_Syntax.role_to_string role ^ "): "), Syntax.pretty_term @{context} fmla])
) (rev an_fmlas)
-*}
+\<close>
-ML {*
+ML \<open>
(*Extract the (name, term) pairs of formulas having roles belonging to a
user-supplied set*)
fun extract_terms roles : TPTP_Interpret.tptp_formula_meaning list ->
@@ -38,9 +38,9 @@
fun role_predicate (_, role, _, _) =
fold (fn r1 => fn b => role = r1 orelse b) roles false
in filter role_predicate #> map (fn (n, _, t, _) => (n, t)) end
-*}
+\<close>
-ML {*
+ML \<open>
(*Use a given tactic on a goal*)
fun prove_conjectures tactic ctxt an_fmlas =
let
@@ -59,15 +59,15 @@
Sledgehammer_Tactics.sledgehammer_with_metis_tac ctxt []
(*FIXME use relevance_override*)
{add = [], del = [], only = false} 1)
-*}
+\<close>
-ML {*
+ML \<open>
@{assert} (is_some (try (auto_prove @{context}) an_fmlas) = false)
-*}
+\<close>
sledgehammer_params [provers = z3, debug]
-ML {*
+ML \<open>
@{assert} (is_some (try (sh_prove @{context}) an_fmlas) = true)
-*}
+\<close>
end
--- a/src/HOL/TPTP/TPTP_Parser_Test.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/TPTP/TPTP_Parser_Test.thy Thu May 26 17:51:22 2016 +0200
@@ -10,7 +10,7 @@
begin
section "Problem-name parsing tests"
-ML {*
+ML \<open>
local
open TPTP_Syntax;
open TPTP_Problem_Name;
@@ -52,27 +52,27 @@
TPTP_Problem_Name.parse_problem_name #>
TPTP_Problem_Name.mangle_problem_name)
(TPTP_Syntax.get_file_list tptp_probs_dir)
-*}
+\<close>
section "Parser tests"
-ML {*
+ML \<open>
TPTP_Parser.parse_expression "" "fof(dt_k4_waybel34, axiom, ~ v3).";
TPTP_Parser.parse_expression "" "thf(dt_k4_waybel34, axiom, ~ ($true | $false)).";
TPTP_Parser.parse_expression ""
"thf(dt_k4_waybel34, axiom, ~ (! [X : $o, Y : ($o > $o)] : ( (X | (Y = Y))))).";
TPTP_Parser.parse_expression "" "tff(dt_k4_waybel34, axiom, ~ ($true)).";
payloads_of it;
-*}
-ML {*
+\<close>
+ML \<open>
TPTP_Parser.parse_expression "" "thf(bla, type, x : $o).";
TPTP_Parser.parse_expression ""
"fof(dt_k4_waybel34, axiom, ~ v3_struct_0(k4_waybel34(A))).";
TPTP_Parser.parse_expression ""
"fof(dt_k4_waybel34, axiom, (! [A] : (v1_xboole_0(A) => ( ~ v3_struct_0(k4_waybel34(A)))))).";
-*}
-ML {*
+\<close>
+ML \<open>
TPTP_Parser.parse_expression ""
("fof(dt_k4_waybel34,axiom,(" ^
"! [A] :" ^
@@ -84,9 +84,9 @@
"& v12_altcat_1(k4_waybel34(A))" ^
"& v2_yellow21(k4_waybel34(A))" ^
"& l2_altcat_1(k4_waybel34(A)) ) ) )).")
-*}
+\<close>
-ML {*
+ML \<open>
open TPTP_Syntax;
@{assert}
((TPTP_Parser.parse_expression ""
@@ -99,30 +99,30 @@
Atom (THF_Atom_term (Term_Func (Uninterpreted "f", []))))),
Atom (THF_Atom_term (Term_Func (Uninterpreted "g", [])))])
)
-*}
+\<close>
text "Parse a specific problem."
-ML {*
+ML \<open>
map TPTP_Parser.parse_file
["$TPTP/Problems/FLD/FLD051-1.p",
"$TPTP/Problems/FLD/FLD005-3.p",
"$TPTP/Problems/SWV/SWV567-1.015.p",
"$TPTP/Problems/SWV/SWV546-1.010.p"]
-*}
-ML {*
+\<close>
+ML \<open>
parser_test @{context} (situate "DAT/DAT001=1.p");
parser_test @{context} (situate "ALG/ALG001^5.p");
parser_test @{context} (situate "NUM/NUM021^1.p");
parser_test @{context} (situate "SYN/SYN000^1.p")
-*}
+\<close>
text "Run the parser over all problems."
-ML {*
+ML \<open>
if test_all @{context} then
(report @{context} "Testing parser";
S timed_test parser_test @{context} (TPTP_Syntax.get_file_list tptp_probs_dir))
else ()
-*}
+\<close>
end
--- a/src/HOL/TPTP/TPTP_Proof_Reconstruction.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/TPTP/TPTP_Proof_Reconstruction.thy Thu May 26 17:51:22 2016 +0200
@@ -52,7 +52,7 @@
section "Setup"
-ML {*
+ML \<open>
val tptp_unexceptional_reconstruction = Attrib.setup_config_bool @{binding tptp_unexceptional_reconstruction} (K false)
fun unexceptional_reconstruction ctxt = Config.get ctxt tptp_unexceptional_reconstruction
val tptp_informative_failure = Attrib.setup_config_bool @{binding tptp_informative_failure} (K false)
@@ -67,7 +67,7 @@
if max = 0 then false
else size > max
end
-*}
+\<close>
(*FIXME move to TPTP_Proof_Reconstruction_Test_Units*)
declare [[
@@ -87,7 +87,7 @@
section "Proof reconstruction"
-text {*There are two parts to proof reconstruction:
+text \<open>There are two parts to proof reconstruction:
\begin{itemize}
\item interpreting the inferences
\item building the skeleton, which indicates how to compose
@@ -144,7 +144,7 @@
clause representation to fit them, and adapting the inference
interpretation to handle the rules used by the prover. It should also
facilitate composing together proofs found by different provers.
-*}
+\<close>
subsection "Instantiation"
@@ -159,7 +159,7 @@
"\<lbrakk>(\<forall>x. P x) = False; \<And>x. (P x) = False \<Longrightarrow> R\<rbrakk> \<Longrightarrow> R"
by auto
-ML {*
+ML \<open>
(*This carries out an allE-like rule but on (polarised) literals.
Instead of yielding a free variable (which is a hell for the
matcher) it seeks to use one of the subgoals' parameters.
@@ -195,9 +195,9 @@
(*Attempts to use the polar_allE theorems on a specific subgoal.*)
fun forall_pos_tac ctxt = inst_parametermatch_tac ctxt @{thms polar_allE}
-*}
+\<close>
-ML {*
+ML \<open>
(*This is similar to inst_parametermatch_tac, but prefers to
match variables having identical names. Logically, this is
a hack. But it reduces the complexity of the problem.*)
@@ -233,12 +233,12 @@
K no_tac i st
end
end
-*}
+\<close>
subsection "Prefix massaging"
-ML {*
+ML \<open>
exception NO_GOALS
(*Get quantifier prefix of the hypothesis and conclusion, reorder
@@ -291,7 +291,7 @@
dresolve_tac ctxt [thm] i st
end
end
-*}
+\<close>
subsection "Some general rules and congruences"
@@ -300,14 +300,14 @@
applied implicitly during some Leo2 inferences.*)
lemma polarise: "P ==> P = True" by auto
-ML {*
+ML \<open>
fun is_polarised t =
(TPTP_Reconstruct.remove_polarity true t; true)
handle TPTP_Reconstruct.UNPOLARISED _ => false
fun polarise_subgoal_hyps ctxt =
COND' (SOME #> TERMPRED is_polarised (fn _ => true)) (K no_tac) (dresolve_tac ctxt @{thms polarise})
-*}
+\<close>
lemma simp_meta [rule_format]:
"(A --> B) == (~A | B)"
@@ -335,12 +335,12 @@
by auto
lemma solved_all_splits: "False = True \<Longrightarrow> False" by simp
-ML {*
+ML \<open>
fun solved_all_splits_tac ctxt =
TRY (eresolve_tac ctxt @{thms conjE} 1)
THEN resolve_tac ctxt @{thms solved_all_splits} 1
THEN assume_tac ctxt 1
-*}
+\<close>
lemma lots_of_logic_expansions_meta [rule_format]:
"(((A :: bool) = B) = True) == (((A \<longrightarrow> B) = True) & ((B \<longrightarrow> A) = True))"
@@ -408,7 +408,7 @@
subsection "Emulation: tactics"
-ML {*
+ML \<open>
(*Instantiate a variable according to the info given in the
proof annotation. Through this we avoid having to come up
with instantiations during reconstruction.*)
@@ -441,9 +441,9 @@
(*now only the variable to instantiate should be left*)
THEN FIRST (map instantiate_tac ordered_instances)
end
-*}
+\<close>
-ML {*
+ML \<open>
(*Simplification tactics*)
local
fun rew_goal_tac thms ctxt i =
@@ -456,7 +456,7 @@
val simper_animal =
rew_goal_tac @{thms simp_meta}
end
-*}
+\<close>
lemma prop_normalise [rule_format]:
"(A | B) | C == A | B | C"
@@ -464,7 +464,7 @@
"A | B == ~(~A & ~B)"
"~~ A == A"
by auto
-ML {*
+ML \<open>
(*i.e., break_conclusion*)
fun flip_conclusion_tac ctxt =
let
@@ -476,7 +476,7 @@
in
default_tac ORELSE' resolve_tac ctxt @{thms flip}
end
-*}
+\<close>
subsection "Skolemisation"
@@ -541,7 +541,7 @@
apply (simp, drule someI_ex, simp)
done
-ML {*
+ML \<open>
(*FIXME LHS should be constant. Currently allow variables for testing. Probably should still allow Vars (but not Frees) since they'll act as intermediate values*)
fun conc_is_skolem_def t =
case t of
@@ -574,9 +574,9 @@
h_property andalso args_property
end
| _ => false
-*}
+\<close>
-ML {*
+ML \<open>
(*Hack used to detect if a Skolem definition, with an LHS Var, has had the LHS instantiated into an unacceptable term.*)
fun conc_is_bad_skolem_def t =
case t of
@@ -610,9 +610,9 @@
h_property andalso args_property
end
| _ => false
-*}
+\<close>
-ML {*
+ML \<open>
fun get_skolem_conc t =
let
val t' =
@@ -633,7 +633,7 @@
|> head_of
|> dest_Const)
(get_skolem_conc t)
-*}
+\<close>
(*
Technique for handling quantifiers:
@@ -643,7 +643,7 @@
or bind a fresh !! -- currently not doing the latter since it never seems to arised in normal Leo2 proofs.
*)
-ML {*
+ML \<open>
fun forall_neg_tac candidate_consts ctxt i = fn st =>
let
val gls =
@@ -678,9 +678,9 @@
(fold (curry (op APPEND')) attempts (K no_tac)) i st
end
end
-*}
+\<close>
-ML {*
+ML \<open>
exception SKOLEM_DEF of term (*The tactic wasn't pointed at a skolem definition*)
exception NO_SKOLEM_DEF of (*skolem const name*)string * Binding.binding * term (*The tactic could not find a skolem definition in the theory*)
fun absorb_skolem_def ctxt prob_name_opt i = fn st =>
@@ -750,9 +750,9 @@
resolve_tac ctxt [Drule.export_without_context thm] i st
end
handle SKOLEM_DEF _ => no_tac st
-*}
+\<close>
-ML {*
+ML \<open>
(*
In current system, there should only be 2 subgoals: the one where
the skolem definition is being built (with a Var in the LHS), and the other subgoal using Var.
@@ -801,9 +801,9 @@
|> maps (get_skolem_terms [] [])
|> distinct (op =)
end
-*}
+\<close>
-ML {*
+ML \<open>
fun instantiate_skols ctxt consts_candidates i = fn st =>
let
val gls =
@@ -953,9 +953,9 @@
in
tactic st
end
-*}
+\<close>
-ML {*
+ML \<open>
fun new_skolem_tac ctxt consts_candidates =
let
fun tac thm =
@@ -965,12 +965,12 @@
if null consts_candidates then K no_tac
else FIRST' (map tac @{thms lift_exists})
end
-*}
+\<close>
(*
need a tactic to expand "? x . P" to "~ ! x. ~ P"
*)
-ML {*
+ML \<open>
fun ex_expander_tac ctxt i =
let
val simpset =
@@ -979,12 +979,12 @@
in
CHANGED (asm_full_simp_tac simpset i)
end
-*}
+\<close>
subsubsection "extuni_dec"
-ML {*
+ML \<open>
(*n-ary decomposition. Code is based on the n-ary arg_cong generator*)
fun extuni_dec_n ctxt arity =
let
@@ -1035,9 +1035,9 @@
Goal.prove ctxt [] [] t (fn _ => auto_tac ctxt)
|> Drule.export_without_context
end
-*}
+\<close>
-ML {*
+ML \<open>
(*Determine the arity of a function which the "dec"
unification rule is about to be applied.
NOTE:
@@ -1157,7 +1157,7 @@
in
(CHANGED o search_tac) i st
end
-*}
+\<close>
subsubsection "standard_cnf"
@@ -1178,7 +1178,7 @@
stage, together with splitting.
*)
-ML {*
+ML \<open>
(*Conjunctive counterparts to Term.disjuncts_aux and Term.disjuncts*)
fun conjuncts_aux (Const (@{const_name HOL.conj}, _) $ t $ t') conjs =
conjuncts_aux t (conjuncts_aux t' conjs)
@@ -1196,9 +1196,9 @@
imp_strip_horn' [] t
|> apfst rev
end
-*}
+\<close>
-ML {*
+ML \<open>
(*Returns whether the antecedents are separated by conjunctions
or implications; the number of antecedents; and the polarity
of the original clause -- I think this will always be "false".*)
@@ -1254,9 +1254,9 @@
if null antes then NONE
else SOME (ante_type, length antes', pol)
end
-*}
+\<close>
-ML {*
+ML \<open>
(*Given a certain standard_cnf type, build a metatheorem that would
validate it*)
fun mk_standard_cnf ctxt kind arity =
@@ -1294,9 +1294,9 @@
Goal.prove ctxt [] [] t (fn _ => HEADGOAL (blast_tac ctxt))
|> Drule.export_without_context
end
-*}
+\<close>
-ML {*
+ML \<open>
(*Applies a d-tactic, then breaks it up conjunctively.
This can be used to transform subgoals as follows:
(A \<longrightarrow> B) = False \<Longrightarrow> R
@@ -1307,15 +1307,15 @@
fun weak_conj_tac ctxt drule =
dresolve_tac ctxt [drule] THEN'
(REPEAT_DETERM o eresolve_tac ctxt @{thms conjE})
-*}
+\<close>
-ML {*
+ML \<open>
fun uncurry_lit_neg_tac ctxt =
REPEAT_DETERM o
dresolve_tac ctxt [@{lemma "(A \<longrightarrow> B \<longrightarrow> C) = False \<Longrightarrow> (A & B \<longrightarrow> C) = False" by auto}]
-*}
+\<close>
-ML {*
+ML \<open>
fun standard_cnf_tac ctxt i = fn st =>
let
fun core_tactic i = fn st =>
@@ -1332,12 +1332,12 @@
THEN' TPTP_Reconstruct_Library.reassociate_conjs_tac ctxt
THEN' core_tactic) i st
end
-*}
+\<close>
subsubsection "Emulator prep"
-ML {*
+ML \<open>
datatype cleanup_feature =
RemoveHypothesesFromSkolemDefs
| RemoveDuplicates
@@ -1380,9 +1380,9 @@
| InnerLoopOnce of loop_feature list
| CleanUp of cleanup_feature list
| AbsorbSkolemDefs
-*}
+\<close>
-ML {*
+ML \<open>
fun can_feature x l =
let
fun sublist_of_clean_up el =
@@ -1442,9 +1442,9 @@
@{assert}
(not (can_feature (Loop []) [InnerLoopOnce [Existential_Var]]));
-*}
+\<close>
-ML {*
+ML \<open>
exception NO_LOOP_FEATS
fun get_loop_feats (feats : feature list) =
let
@@ -1467,11 +1467,11 @@
@{assert}
(get_loop_feats [Loop [King_Cong, Break_Hypotheses, Existential_Free, Existential_Var, Universal]] =
[King_Cong, Break_Hypotheses, Existential_Free, Existential_Var, Universal])
-*}
+\<close>
(*use as elim rule to remove premises*)
lemma insa_prems: "\<lbrakk>Q; P\<rbrakk> \<Longrightarrow> P" by auto
-ML {*
+ML \<open>
fun cleanup_skolem_defs ctxt feats =
let
(*remove hypotheses from skolem defs,
@@ -1485,16 +1485,16 @@
ALLGOALS (TRY o dehypothesise_skolem_defs)
else all_tac
end
-*}
+\<close>
-ML {*
+ML \<open>
fun remove_duplicates_tac feats =
(if can_feature (CleanUp [RemoveDuplicates]) feats then
ALLGOALS distinct_subgoal_tac
else all_tac)
-*}
+\<close>
-ML {*
+ML \<open>
(*given a goal state, indicates the skolem constants committed-to in it (i.e. appearing in LHS of a skolem definition)*)
fun which_skolem_concs_used ctxt = fn st =>
let
@@ -1512,9 +1512,9 @@
|> switch (fold (fn x => fn l => if is_some x then the x :: l else l)) []
|> map Const
end
-*}
+\<close>
-ML {*
+ML \<open>
fun exists_tac ctxt feats consts_diff =
let
val ex_var =
@@ -1535,7 +1535,7 @@
if loop_can_feature [Universal] feats then
forall_pos_tac ctxt
else K no_tac
-*}
+\<close>
subsubsection "Finite types"
@@ -1545,7 +1545,7 @@
(*predicate over the type of the leading quantified variable*)
-ML {*
+ML \<open>
fun extcnf_forall_special_pos_tac ctxt =
let
val bool =
@@ -1565,19 +1565,19 @@
(*FIXME could check the type of the leading quantified variable, instead of trying everything*)
(tacs (bool @ bool_to_bool)))
end
-*}
+\<close>
subsubsection "Emulator"
lemma efq: "[|A = True; A = False|] ==> R" by auto
-ML {*
+ML \<open>
fun efq_tac ctxt =
(eresolve_tac ctxt @{thms efq} THEN' assume_tac ctxt)
ORELSE' assume_tac ctxt
-*}
+\<close>
-ML {*
+ML \<open>
(*This is applied to all subgoals, repeatedly*)
fun extcnf_combined_main ctxt feats consts_diff =
let
@@ -1697,9 +1697,9 @@
if head_of t = Logic.implies then do_diff t
else []
end
-*}
+\<close>
-ML {*
+ML \<open>
(*remove quantification in hypothesis clause (! X. t), if
X not free in t*)
fun remove_redundant_quantification ctxt i = fn st =>
@@ -1740,14 +1740,14 @@
end
end
end
-*}
+\<close>
-ML {*
+ML \<open>
fun remove_redundant_quantification_ignore_skolems ctxt i =
COND (TERMPRED (fn _ => true) conc_is_skolem_def (SOME i))
no_tac
(remove_redundant_quantification ctxt i)
-*}
+\<close>
lemma drop_redundant_literal_qtfr:
"(! X. P) = True \<Longrightarrow> P = True"
@@ -1756,7 +1756,7 @@
"(? X. P) = False \<Longrightarrow> P = False"
by auto
-ML {*
+ML \<open>
(*remove quantification in the literal "(! X. t) = True/False"
in the singleton hypothesis clause, if X not free in t*)
fun remove_redundant_quantification_in_lit ctxt i = fn st =>
@@ -1808,16 +1808,16 @@
end
end
end
-*}
+\<close>
-ML {*
+ML \<open>
fun remove_redundant_quantification_in_lit_ignore_skolems ctxt i =
COND (TERMPRED (fn _ => true) conc_is_skolem_def (SOME i))
no_tac
(remove_redundant_quantification_in_lit ctxt i)
-*}
+\<close>
-ML {*
+ML \<open>
fun extcnf_combined_tac ctxt prob_name_opt feats skolem_consts = fn st =>
let
val thy = Proof_Context.theory_of ctxt
@@ -1889,7 +1889,7 @@
in
DEPTH_SOLVE (CHANGED tec) st
end
-*}
+\<close>
subsubsection "unfold_def"
@@ -1904,7 +1904,7 @@
lemma un_meta_polarise: "(X \<equiv> True) \<Longrightarrow> X" by auto
lemma meta_polarise: "X \<Longrightarrow> X \<equiv> True" by auto
-ML {*
+ML \<open>
fun unfold_def_tac ctxt depends_on_defs = fn st =>
let
(*This is used when we end up with something like
@@ -1946,7 +1946,7 @@
in
tactic st
end
-*}
+\<close>
subsection "Handling split 'preprocessing'"
@@ -1960,7 +1960,7 @@
by (rule eq_reflection, auto)+
(*Same idiom as ex_expander_tac*)
-ML {*
+ML \<open>
fun split_simp_tac (ctxt : Proof.context) i =
let
val simpset =
@@ -1968,11 +1968,11 @@
in
CHANGED (asm_full_simp_tac simpset i)
end
-*}
+\<close>
subsection "Alternative reconstruction tactics"
-ML {*
+ML \<open>
(*An "auto"-based proof reconstruction, where we attempt to reconstruct each inference
using auto_tac. A realistic tactic would inspect the inference name and act
accordingly.*)
@@ -1989,11 +1989,11 @@
Goal.prove ctxt [] [] inference_fmla
(fn pdata => auto_tac (#context pdata)))
end
-*}
+\<close>
(*An oracle-based reconstruction, which is only used to test the shunting part of the system*)
oracle oracle_iinterp = "fn t => t"
-ML {*
+ML \<open>
fun oracle_based_reconstruction_tac ctxt prob_name n =
let
val thy = Proof_Context.theory_of ctxt
@@ -2006,12 +2006,12 @@
|> (fn {inference_fmla, ...} => Thm.cterm_of ctxt inference_fmla)
|> oracle_iinterp
end
-*}
+\<close>
subsection "Leo2 reconstruction tactic"
-ML {*
+ML \<open>
exception UNSUPPORTED_ROLE
exception INTERPRET_INFERENCE
@@ -2171,9 +2171,9 @@
| "flexflex" => default_tac
| other => fail ctxt ("Unknown inference rule: " ^ other)
end
-*}
+\<close>
-ML {*
+ML \<open>
fun interpret_leo2_inference ctxt prob_name node =
let
val thy = Proof_Context.theory_of ctxt
@@ -2204,9 +2204,9 @@
Syntax.pretty_term ctxt inference_fmla]))
| SOME thm => thm
end
-*}
+\<close>
-ML {*
+ML \<open>
(*filter a set of nodes based on which inference rule was used to
derive a node*)
fun nodes_by_inference (fms : TPTP_Reconstruct.formula_meaning list) inference_rule =
@@ -2221,12 +2221,12 @@
in
fold fold_fun (map fst fms) []
end
-*}
+\<close>
section "Importing proofs and reconstructing theorems"
-ML {*
+ML \<open>
(*Preprocessing carried out on a LEO-II proof.*)
fun leo2_on_load (pannot : TPTP_Reconstruct.proof_annotation) thy =
let
@@ -2256,9 +2256,9 @@
meta = #meta pannot},
thy')
end
-*}
+\<close>
-ML {*
+ML \<open>
(*Imports and reconstructs a LEO-II proof.*)
fun reconstruct_leo2 path thy =
let
@@ -2280,6 +2280,6 @@
users could get the thy value from the thm value.*)
(thy', theorem)
end
-*}
+\<close>
end
--- a/src/HOL/TPTP/TPTP_Proof_Reconstruction_Test.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/TPTP/TPTP_Proof_Reconstruction_Test.thy Thu May 26 17:51:22 2016 +0200
@@ -23,9 +23,9 @@
proof have been imported), but note that you can get the theory value from
the theorem value."
-ML {*
+ML \<open>
reconstruct_leo2 (Path.explode "$THF_PROOFS/NUM667^1.p.out") @{theory}
-*}
+\<close>
section "Prep for testing the component functions"
@@ -43,7 +43,7 @@
section "Importing proofs"
-ML {*
+ML \<open>
val probs =
(* "$THF_PROOFS/SYN991^1.p.out" *) (*lacks conjecture*)
(* "$THF_PROOFS/SYO040^2.p.out" *)
@@ -59,19 +59,19 @@
val prob_names =
probs
|> map (Path.base #> Path.implode #> TPTP_Problem_Name.Nonstandard)
-*}
+\<close>
-setup {*
+setup \<open>
if test_all @{context} then I
else
fold
(fn path =>
TPTP_Reconstruct.import_thm true [Path.dir path, Path.explode "$THF_PROOFS"] path leo2_on_load)
probs
-*}
+\<close>
text "Display nicely."
-ML {*
+ML \<open>
fun display_nicely ctxt (fms : TPTP_Reconstruct.formula_meaning list) =
List.app (fn ((n, data) : TPTP_Reconstruct.formula_meaning) =>
Pretty.writeln
@@ -100,9 +100,9 @@
(test_fmla @{theory}
|> map TPTP_Reconstruct.structure_fmla_meaning)
*)
-*}
+\<close>
-ML {*
+ML \<open>
fun step_range_tester f_x f_exn ctxt prob_name from until =
let
val max =
@@ -126,9 +126,9 @@
step_range_tester
(fn x => tracing ("@step " ^ Int.toString x))
(fn e => tracing ("!!" ^ @{make_string} e))
-*}
+\<close>
-ML {*
+ML \<open>
(*try to reconstruct each inference step*)
if test_all @{context} orelse prob_names = []
orelse true (*NOTE currently disabled*)
@@ -144,48 +144,48 @@
in
step_range_tester_tracing @{context} (hd prob_names) heur_start heur_end
end
-*}
+\<close>
section "Building metadata and tactics"
subsection "Building the skeleton"
-ML {*
+ML \<open>
if test_all @{context} orelse prob_names = [] then []
else TPTP_Reconstruct.make_skeleton @{context} (test_pannot @{theory});
length it
-*}
+\<close>
subsection "The 'one shot' tactic approach"
-ML {*
+ML \<open>
val the_tactic =
if test_all @{context} then []
else
map (fn prob_name =>
(TPTP_Reconstruct.naive_reconstruct_tac @{context} interpret_leo2_inference (* auto_based_reconstruction_tac *) (* oracle_based_reconstruction_tac *) prob_name))
prob_names;
-*}
+\<close>
subsection "The 'piecemeal' approach"
-ML {*
+ML \<open>
val the_tactics =
if test_all @{context} then []
else
map (fn prob_name =>
TPTP_Reconstruct.naive_reconstruct_tacs interpret_leo2_inference (* auto_based_reconstruction_tac *) (* oracle_based_reconstruction_tac *) prob_name @{context})
prob_names;
-*}
+\<close>
declare [[ML_print_depth = 2000]]
-ML {*
+ML \<open>
the_tactics
|> map (filter (fn (_, _, x) => is_none x)
#> map (fn (x, SOME y, _) => (x, cterm_of @{theory} y)))
-*}
+\<close>
section "Using metadata and tactics"
@@ -194,7 +194,7 @@
subsection "The 'one shot' tactic approach"
text "First we test whole tactics."
-ML {*
+ML \<open>
(*produce thm*)
if test_all @{context} then []
else
@@ -204,18 +204,18 @@
TPTP_Reconstruct.naive_reconstruct_tac @{context} interpret_leo2_inference prob_name
(* oracle_based_reconstruction_tac *))))
prob_names
-*}
+\<close>
subsection "The 'piecemeal' approach"
-ML {*
+ML \<open>
fun attac n = List.nth (List.nth (the_tactics, 0), n) |> #3 |> the |> snd
fun attac_to n 0 = attac n
| attac_to n m = attac n THEN attac_to (n + 1) (m - 1)
fun shotac n = List.nth (List.nth (the_tactics, 0), n) |> #3 |> the |> fst
-*}
+\<close>
-ML {*
+ML \<open>
(*Given a list of reconstructed inferences (as in "the_tactics" above,
count the number of failures and successes, and list the failed inference
reconstructions.*)
@@ -234,11 +234,11 @@
end
| evaluate_the_tactics ((_, (_, _, SOME _)) :: xs) ((fai, suc), inf_list) =
evaluate_the_tactics xs ((fai, suc + 1), inf_list)
-*}
+\<close>
text "Now we build a tactic by combining lists of tactics"
-ML {*
+ML \<open>
(*given a list of tactics to be applied in sequence (i.e., they
follow a skeleton), we build a single tactic, interleaving
some tracing info to help with debugging.*)
@@ -261,9 +261,9 @@
|> uncurry interleave_tacs
else EVERY (map #2 thm_tacs)
end
-*}
+\<close>
-ML {*
+ML \<open>
(*apply step_by_step_tacs to all problems under test*)
fun narrated_tactics ctxt =
map (map (#3 #> the)
@@ -278,13 +278,13 @@
TPTP_Reconstruct.reconstruct @{context}
(fn _ => tac) prob_name)
(ListPair.zip (prob_names, narrated_tactics @{context}))
-*}
+\<close>
subsection "Manually using 'piecemeal' approach"
text "Another testing possibility involves manually creating a lemma
and running through the list of tactics generating to prove that lemma. The following code shows the goal of each problem under test, and then for each problem returns the list of tactics which can be invoked individually as shown below."
-ML {*
+ML \<open>
fun show_goal ctxt prob_name =
let
val thy = Proof_Context.theory_of ctxt
@@ -301,36 +301,36 @@
if test_all @{context} then []
else
map (show_goal @{context}) prob_names;
-*}
+\<close>
-ML {*
+ML \<open>
(*project out the list of tactics from "the_tactics"*)
val just_the_tacs =
map (map (#3 #> the #> #2))
the_tactics;
map length just_the_tacs
-*}
+\<close>
-ML {*
+ML \<open>
(*like just_the_tacs, but extract the thms, to inspect their thys*)
val just_the_thms =
map (map (#3 #> the #> #1))
the_tactics;
map length just_the_thms;
-*}
+\<close>
-ML {*
+ML \<open>
val axms_of_thy =
`Theory.axioms_of
#> apsnd cterm_of
#> swap
#> apsnd (map snd)
#> uncurry map
-*}
+\<close>
-ML {*
+ML \<open>
(*Show the skeleton-level inference which is done by each element of just_the_tacs. This is useful when debugging using the technique shown next*)
if test_all @{context} orelse prob_names = [] then ()
else
@@ -339,12 +339,12 @@
|> map #1
|> TPTP_Reconstruct_Library.enumerate 0
|> List.app (@{make_string} #> writeln)
- *}
+\<close>
-ML {*
+ML \<open>
fun leo2_tac_wrap ctxt prob_name step i st =
rtac (interpret_leo2_inference ctxt prob_name step) i st
-*}
+\<close>
(*FIXME move these examples elsewhere*)
(*
@@ -552,7 +552,7 @@
section "Testing against benchmark"
-ML {*
+ML \<open>
(*if reconstruction_info value is NONE then a big error must have occurred*)
type reconstruction_info =
((int(*no of failures*) * int(*no of successes*)) *
@@ -567,9 +567,9 @@
fun erase_inference_fmlas (Nonempty (SOME (outline, inf_info))) =
Nonempty (SOME (outline, map (fn (inf_name, _, count) => (inf_name, NONE, count)) inf_info))
| erase_inference_fmlas x = x
-*}
+\<close>
-ML {*
+ML \<open>
(*Report on how many inferences in a proof are reconstructed, and give some
info about the inferences for which reconstruction failed.*)
fun test_partial_reconstruction thy prob_file =
@@ -626,9 +626,9 @@
if is_some fms andalso List.null (the fms) then Empty
else Nonempty result'
end
-*}
+\<close>
-ML {*
+ML \<open>
(*default timeout is 1 min*)
fun reconstruct timeout light_output file thy =
let
@@ -642,9 +642,9 @@
" t=" ^ (Timer.checkRealTimer timer |> Time.toMilliseconds |> @{make_string}))))
file
end
-*}
+\<close>
-ML {*
+ML \<open>
(*this version of "reconstruct" builds theorems, instead of lists of reconstructed inferences*)
(*default timeout is 1 min*)
fun reconstruct timeout file thy =
@@ -671,17 +671,17 @@
" t=" ^ (Timer.checkRealTimer timer |> Time.toMilliseconds |> @{make_string}))))
prob_name
end
-*}
+\<close>
-ML {*
+ML \<open>
fun reconstruction_test timeout ctxt =
test_fn ctxt
(fn file => reconstruct timeout file (Proof_Context.theory_of ctxt))
"reconstructor"
()
-*}
+\<close>
-ML {*
+ML \<open>
datatype examination_results =
Whole_proof of string(*filename*) * proof_contents
| Specific_rule of string(*filename*) * string(*inference rule*) * term option list
@@ -713,9 +713,9 @@
|> is_some rule_name ? (fn x =>
filter_failures (the rule_name) x)
end
-*}
+\<close>
-ML {*
+ML \<open>
exception NONSENSE
fun annotation_or_id (TPTP_Reconstruct.Step n) = n
@@ -748,9 +748,9 @@
fun classify_failures (Whole_proof (_, Nonempty (SOME (((_, _), inferences))))) = pre_classify_failures inferences []
| classify_failures (Specific_rule (_, rule, t)) = [(rule, length t)]
| classify_failures _ = raise NONSENSE
-*}
+\<close>
-ML {*
+ML \<open>
val regressions = map (fn s => "$THF_PROOFS/" ^ s)
["SEV405^5.p.out",
(*"SYO377^5.p.out", Always seems to raise Interrupt on my laptop -- probably because node 475 has lots of premises*)
@@ -775,9 +775,9 @@
"SYO006^1.p.out",
"SYO371^5.p.out" (*has contorted splitting, like SYO006^1.p.out, but doesn't involve definitions*)
]
-*}
+\<close>
-ML {*
+ML \<open>
val experiment = examine_failed_inferences @{context}
(List.last regressions) NONE;
@@ -790,7 +790,7 @@
count_failures experiment_focus
classify_failures experiment
*)
-*}
+\<close>
text "Run reconstruction on all problems in a benchmark (provided via a script)
and report on partial success."
@@ -800,19 +800,19 @@
tptp_test_timeout = 10
]]
-ML {*
+ML \<open>
(*problem source*)
val tptp_probs_dir =
Path.explode "$THF_PROOFS"
|> Path.expand;
-*}
+\<close>
-ML {*
+ML \<open>
if test_all @{context} then
(report @{context} "Reconstructing proofs";
S timed_test (reconstruction_test (get_timeout @{context})) @{context} (TPTP_Syntax.get_file_list tptp_probs_dir))
else ()
-*}
+\<close>
(*
Debugging strategy:
--- a/src/HOL/TPTP/TPTP_Proof_Reconstruction_Test_Units.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/TPTP/TPTP_Proof_Reconstruction_Test_Units.thy Thu May 26 17:51:22 2016 +0200
@@ -13,13 +13,13 @@
declare [[tptp_trace_reconstruction = true]]
lemma "! (X1 :: bool) (X2 :: bool) (X3 :: bool) (X4 :: bool) (X5 :: bool). P \<Longrightarrow> ! (X1 :: bool) (X3 :: bool) (X5 :: bool). P"
-apply (tactic {*canonicalise_qtfr_order @{context} 1*})
+apply (tactic \<open>canonicalise_qtfr_order @{context} 1\<close>)
oops
lemma "! (X1 :: bool) (X2 :: bool) (X3 :: bool) (X4 :: bool) (X5 :: bool). P \<Longrightarrow> ! (X1 :: bool) (X3 :: bool) (X5 :: bool). P"
-apply (tactic {*canonicalise_qtfr_order @{context} 1*})
+apply (tactic \<open>canonicalise_qtfr_order @{context} 1\<close>)
apply (rule allI)+
-apply (tactic {*nominal_inst_parametermatch_tac @{context} @{thm allE} 1*})+
+apply (tactic \<open>nominal_inst_parametermatch_tac @{context} @{thm allE} 1\<close>)+
oops
(*Could test bind_tac further with NUM667^1 inode43*)
@@ -187,35 +187,35 @@
lemma "(A = B) = True ==> (B = A) = True"
*)
lemma "!!x. ((A x :: bool) = B x) = False ==> (B x = A x) = False"
-apply (tactic {*expander_animal @{context} 1*})
+apply (tactic \<open>expander_animal @{context} 1\<close>)
oops
lemma "(A & B) ==> ~(~A | ~B)"
-by (tactic {*expander_animal @{context} 1*})
+by (tactic \<open>expander_animal @{context} 1\<close>)
lemma "(A | B) ==> ~(~A & ~B)"
-by (tactic {*expander_animal @{context} 1*})
+by (tactic \<open>expander_animal @{context} 1\<close>)
lemma "(A | B) | C ==> A | (B | C)"
-by (tactic {*expander_animal @{context} 1*})
+by (tactic \<open>expander_animal @{context} 1\<close>)
lemma "(~~A) = B ==> A = B"
-by (tactic {*expander_animal @{context} 1*})
+by (tactic \<open>expander_animal @{context} 1\<close>)
lemma "~ ~ (A = True) ==> A = True"
-by (tactic {*expander_animal @{context} 1*})
+by (tactic \<open>expander_animal @{context} 1\<close>)
(*This might not be a goal which might realistically arise:
lemma "((~~A) = B) & (B = (~~A)) ==> ~(~(A = B) | ~(B = A))" *)
lemma "((~~A) = True) ==> ~(~(A = True) | ~(True = A))"
-apply (tactic {*expander_animal @{context} 1*})+
+apply (tactic \<open>expander_animal @{context} 1\<close>)+
apply (rule conjI)
apply assumption
apply (rule sym, assumption)
done
lemma "A = B ==> ((~~A) = B) & (B = (~~A)) ==> ~(~(A = B) | ~(B = A))"
-by (tactic {*expander_animal @{context} 1*})+
+by (tactic \<open>expander_animal @{context} 1\<close>)+
(*some lemmas assume constants in the signature of PUZ114^5*)
consts
@@ -242,7 +242,7 @@
SY30 (PUZ114_5_bnd_s Xj) (PUZ114_5_bnd_s Xk)))
| SY30 PUZ114_5_bnd_sK1 PUZ114_5_bnd_sK2
)"
-by (tactic {*expander_animal @{context} 1*})+
+by (tactic \<open>expander_animal @{context} 1\<close>)+
(*
extcnf_forall_pos:
@@ -262,8 +262,8 @@
(where X' is fresh, or renamings are done suitably).*)
lemma "A | B \<Longrightarrow> A | B | C"
-apply (tactic {*flip_conclusion_tac @{context} 1*})+
-apply (tactic {*break_hypotheses 1*})+
+apply (tactic \<open>flip_conclusion_tac @{context} 1\<close>)+
+apply (tactic \<open>break_hypotheses 1\<close>)+
oops
consts
@@ -297,7 +297,7 @@
True) =
False"
apply (rule allI, erule_tac x = "SV2" in allE)
-apply (tactic {*extuni_dec_tac @{context} 1*})
+apply (tactic \<open>extuni_dec_tac @{context} 1\<close>)
done
(*SEU882^5*)
@@ -330,12 +330,12 @@
lemma "A | B \<Longrightarrow> C1 | A | C2 | B | C3"
apply (erule disjE)
-apply (tactic {*clause_breaker 1*})
-apply (tactic {*clause_breaker 1*})
+apply (tactic \<open>clause_breaker 1\<close>)
+apply (tactic \<open>clause_breaker 1\<close>)
done
lemma "A \<Longrightarrow> A"
-apply (tactic {*clause_breaker 1*})
+apply (tactic \<open>clause_breaker 1\<close>)
done
typedecl NUM667_1_bnd_nat
@@ -359,15 +359,15 @@
NUM667_1_bnd_less SV9 SV10 = True) \<or>
(SV9 = SV11) =
False"
-apply (tactic {*strip_qtfrs_tac @{context}*})
-apply (tactic {*break_hypotheses 1*})
-apply (tactic {*ALLGOALS (TRY o clause_breaker)*})
-apply (tactic {*extuni_dec_tac @{context} 1*})
+apply (tactic \<open>strip_qtfrs_tac @{context}\<close>)
+apply (tactic \<open>break_hypotheses 1\<close>)
+apply (tactic \<open>ALLGOALS (TRY o clause_breaker)\<close>)
+apply (tactic \<open>extuni_dec_tac @{context} 1\<close>)
done
-ML {*
+ML \<open>
extuni_dec_n @{context} 2;
-*}
+\<close>
(*NUM667^1, node 202*)
lemma "\<forall>SV9 SV10 SV11.
@@ -382,10 +382,10 @@
NUM667_1_bnd_less SV9 SV10 = True) \<or>
NUM667_1_bnd_less NUM667_1_bnd_x NUM667_1_bnd_y =
True"
-apply (tactic {*strip_qtfrs_tac @{context}*})
-apply (tactic {*break_hypotheses 1*})
-apply (tactic {*ALLGOALS (TRY o clause_breaker)*})
-apply (tactic {*extuni_dec_tac @{context} 1*})
+apply (tactic \<open>strip_qtfrs_tac @{context}\<close>)
+apply (tactic \<open>break_hypotheses 1\<close>)
+apply (tactic \<open>ALLGOALS (TRY o clause_breaker)\<close>)
+apply (tactic \<open>extuni_dec_tac @{context} 1\<close>)
done
(*NUM667^1 node 141*)
@@ -400,7 +400,7 @@
done
*)
-ML {*
+ML \<open>
fun full_extcnf_combined_tac ctxt =
extcnf_combined_tac ctxt NONE
[ConstsDiff,
@@ -418,9 +418,9 @@
CleanUp [RemoveHypothesesFromSkolemDefs, RemoveDuplicates],
AbsorbSkolemDefs]
[]
-*}
+\<close>
-ML {*
+ML \<open>
fun nonfull_extcnf_combined_tac ctxt feats =
extcnf_combined_tac ctxt NONE
[ConstsDiff,
@@ -428,7 +428,7 @@
InnerLoopOnce (Break_Hypotheses :: feats),
AbsorbSkolemDefs]
[]
-*}
+\<close>
consts SEU882_5_bnd_sK1_Xy :: TPTP_Interpret.ind
lemma
@@ -437,7 +437,7 @@
(* apply (erule_tac x = "(@X. False)" in allE) *)
(* apply (tactic {*remove_redundant_quantification 1*}) *)
(* apply assumption *)
-by (tactic {*nonfull_extcnf_combined_tac @{context} [RemoveRedundantQuantifications, Extuni_FlexRigid]*})
+by (tactic \<open>nonfull_extcnf_combined_tac @{context} [RemoveRedundantQuantifications, Extuni_FlexRigid]\<close>)
(*NUM667^1*)
(*
@@ -472,7 +472,7 @@
(\<not> SEU602_2_bnd_in (SEU602_2_bnd_sK7_E SV49) SEU602_2_bnd_sK2_SY17)) =
False"
(*FIXME this (and similar) tests are getting the "Bad background theory of goal state" error since upgrading to Isabelle2013-2.*)
-by (tactic {*nonfull_extcnf_combined_tac @{context} [Extuni_Func, Existential_Var]*})
+by (tactic \<open>nonfull_extcnf_combined_tac @{context} [Extuni_Func, Existential_Var]\<close>)
consts
SEV405_5_bnd_sK1_SY2 :: "(TPTP_Interpret.ind \<Rightarrow> bool) \<Rightarrow> TPTP_Interpret.ind"
@@ -487,7 +487,7 @@
(\<not> (\<not> (\<not> SV1 (SEV405_5_bnd_sK1_SY2 SV1) \<or> SEV405_5_bnd_cA) \<or>
\<not> (\<not> SEV405_5_bnd_cA \<or> SV1 (SEV405_5_bnd_sK1_SY2 SV1)))) =
False"
-by (tactic {*nonfull_extcnf_combined_tac @{context} [Existential_Var]*})
+by (tactic \<open>nonfull_extcnf_combined_tac @{context} [Existential_Var]\<close>)
(*
strip quantifiers -- creating a space of permutations; from shallowest to deepest (iterative deepening)
flip the conclusion -- giving us branch
@@ -503,38 +503,38 @@
(*testing parameters*)
lemma "! X :: TPTP_Interpret.ind . (\<forall>A B. SEU581_2_bnd_in B (SEU581_2_bnd_powerset A) \<longrightarrow> SEU581_2_bnd_subset B A) = True
\<Longrightarrow> ! X :: TPTP_Interpret.ind . (\<forall>A B. \<not> SEU581_2_bnd_in B (SEU581_2_bnd_powerset A) \<or> SEU581_2_bnd_subset B A) = True"
-by (tactic {*full_extcnf_combined_tac @{context}*})
+by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
lemma "(A & B) = True ==> (A | B) = True"
-by (tactic {*full_extcnf_combined_tac @{context}*})
+by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
lemma "(\<not> bnd_subset (bnd_dsetconstr bnd_sK1 bnd_sK2) bnd_sK1) = True \<Longrightarrow> (bnd_subset (bnd_dsetconstr bnd_sK1 bnd_sK2) bnd_sK1) = False"
-by (tactic {*full_extcnf_combined_tac @{context}*})
+by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
(*testing goals with parameters*)
lemma "(\<not> bnd_subset (bnd_dsetconstr bnd_sK1 bnd_sK2) bnd_sK1) = True \<Longrightarrow> ! X. (bnd_subset (bnd_dsetconstr bnd_sK1 bnd_sK2) bnd_sK1) = False"
-by (tactic {*full_extcnf_combined_tac @{context}*})
+by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
lemma "(A & B) = True ==> (B & A) = True"
-by (tactic {*full_extcnf_combined_tac @{context}*})
+by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
(*appreciating differences between THEN, REPEAT, and APPEND*)
lemma "A & B ==> B & A"
-apply (tactic {*
+apply (tactic \<open>
TRY (etac @{thm conjE} 1)
- THEN TRY (rtac @{thm conjI} 1)*})
+ THEN TRY (rtac @{thm conjI} 1)\<close>)
by assumption+
lemma "A & B ==> B & A"
-by (tactic {*
+by (tactic \<open>
etac @{thm conjE} 1
THEN rtac @{thm conjI} 1
- THEN REPEAT (atac 1)*})
+ THEN REPEAT (atac 1)\<close>)
lemma "A & B ==> B & A"
-apply (tactic {*
+apply (tactic \<open>
rtac @{thm conjI} 1
- APPEND etac @{thm conjE} 1*})+
+ APPEND etac @{thm conjE} 1\<close>)+
back
by assumption+
@@ -564,16 +564,16 @@
\<not> SV1 bnd_c1 bnd_c1) \<or>
SV1 PUZ114_5_bnd_sK1 PUZ114_5_bnd_sK2) =
True"
-by (tactic {*full_extcnf_combined_tac @{context}*})
+by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
lemma "(\<not> SEU581_2_bnd_subset (SEU581_2_bnd_dsetconstr SEU581_2_bnd_sK1 SEU581_2_bnd_sK2) SEU581_2_bnd_sK1) = True \<Longrightarrow> (SEU581_2_bnd_subset (SEU581_2_bnd_dsetconstr SEU581_2_bnd_sK1 SEU581_2_bnd_sK2) SEU581_2_bnd_sK1) = False"
-by (tactic {*full_extcnf_combined_tac @{context}*})
+by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
(*testing repeated application of simulator*)
lemma "(\<not> \<not> False) = True \<Longrightarrow>
SEU581_2_bnd_subset (SEU581_2_bnd_dsetconstr SEU581_2_bnd_sK1 SEU581_2_bnd_sK2) SEU581_2_bnd_sK1 = True \<or>
False = True \<or> False = True \<or> False = True"
-by (tactic {*full_extcnf_combined_tac @{context}*})
+by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
(*Testing non-normal conclusion. Ideally we should be able to apply
the tactic to arbitrary chains of extcnf steps -- where it's not
@@ -581,48 +581,48 @@
lemma "(\<not> \<not> False) = True \<Longrightarrow>
SEU581_2_bnd_subset (SEU581_2_bnd_dsetconstr SEU581_2_bnd_sK1 SEU581_2_bnd_sK2) SEU581_2_bnd_sK1 = True \<or>
(\<not> False) = False"
-by (tactic {*full_extcnf_combined_tac @{context}*})
+by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
(*testing repeated application of simulator, involving different extcnf rules*)
lemma "(\<not> \<not> (False | False)) = True \<Longrightarrow>
SEU581_2_bnd_subset (SEU581_2_bnd_dsetconstr SEU581_2_bnd_sK1 SEU581_2_bnd_sK2) SEU581_2_bnd_sK1 = True \<or>
False = True \<or> False = True \<or> False = True"
-by (tactic {*full_extcnf_combined_tac @{context}*})
+by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
(*testing logical expansion*)
lemma "(\<forall>A B. SEU581_2_bnd_in B (SEU581_2_bnd_powerset A) \<longrightarrow> SEU581_2_bnd_subset B A) = True
\<Longrightarrow> (\<forall>A B. \<not> SEU581_2_bnd_in B (SEU581_2_bnd_powerset A) \<or> SEU581_2_bnd_subset B A) = True"
-by (tactic {*full_extcnf_combined_tac @{context}*})
+by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
(*testing extcnf_forall_pos*)
lemma "(\<forall>A Xphi. SEU581_2_bnd_in (SEU581_2_bnd_dsetconstr A Xphi) (SEU581_2_bnd_powerset A)) = True \<Longrightarrow> \<forall>SV1. (\<forall>SY14.
SEU581_2_bnd_in (SEU581_2_bnd_dsetconstr SV1 SY14)
(SEU581_2_bnd_powerset SV1)) = True"
-by (tactic {*full_extcnf_combined_tac @{context}*})
+by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
lemma "((\<forall>A Xphi. SEU581_2_bnd_in (SEU581_2_bnd_dsetconstr A Xphi) (SEU581_2_bnd_powerset A)) = True) | ((~ False) = False) \<Longrightarrow>
\<forall>SV1. ((\<forall>SY14. SEU581_2_bnd_in (SEU581_2_bnd_dsetconstr SV1 SY14) (SEU581_2_bnd_powerset SV1)) = True) | ((~ False) = False)"
-by (tactic {*full_extcnf_combined_tac @{context}*})
+by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
(*testing parameters*)
lemma "(\<forall>A B. SEU581_2_bnd_in B (SEU581_2_bnd_powerset A) \<longrightarrow> SEU581_2_bnd_subset B A) = True
\<Longrightarrow> ! X. (\<forall>A B. \<not> SEU581_2_bnd_in B (SEU581_2_bnd_powerset A) \<or> SEU581_2_bnd_subset B A) = True"
-by (tactic {*full_extcnf_combined_tac @{context}*})
+by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
lemma "((? A .P1 A) = False) | P2 = True \<Longrightarrow> !X. ((P1 X) = False | P2 = True)"
-by (tactic {*full_extcnf_combined_tac @{context}*})
+by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
lemma "((!A . (P1a A | P1b A)) = True) | (P2 = True) \<Longrightarrow> !X. (P1a X = True | P1b X = True | P2 = True)"
-by (tactic {*full_extcnf_combined_tac @{context}*})
+by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
lemma "! Y. (((!A .(P1a A | P1b A)) = True) | P2 = True) \<Longrightarrow> ! Y. (!X. (P1a X = True | P1b X = True | P2 = True))"
-by (tactic {*full_extcnf_combined_tac @{context}*})
+by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
lemma "! Y. (((!A .(P1a A | P1b A)) = True) | P2 = True) \<Longrightarrow> ! Y. (!X. (P1a X = True | P1b X = True | P2 = True))"
-by (tactic {*full_extcnf_combined_tac @{context}*})
+by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
lemma "! Y. (((!A .(P1a A | P1b A)) = True) | P2 = True) \<Longrightarrow> ! Y. (!X. (P1a X = True | P1b X = True | P2 = True))"
-by (tactic {*full_extcnf_combined_tac @{context}*})
+by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
consts dud_bnd_s :: "TPTP_Interpret.ind \<Rightarrow> TPTP_Interpret.ind"
@@ -636,12 +636,12 @@
\<not> PUZ114_5_bnd_sK3 SX0 SX1 \<or>
PUZ114_5_bnd_sK3 (dud_bnd_s SX0) (dud_bnd_s SX1))) =
False"
-by (tactic {*full_extcnf_combined_tac @{context}*})
+by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
(*testing logical expansion -- this should be done by blast*)
lemma "(\<forall>A B. bnd_in B (bnd_powerset A) \<longrightarrow> SEU581_2_bnd_subset B A) = True
\<Longrightarrow> (\<forall>A B. \<not> bnd_in B (bnd_powerset A) \<or> SEU581_2_bnd_subset B A) = True"
-by (tactic {*full_extcnf_combined_tac @{context}*})
+by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
(*testing related to PUZ114^5.p.out*)
lemma "\<forall>SV1. ((\<not> (\<not> SV1 (PUZ114_5_bnd_sK4 SV1) (PUZ114_5_bnd_sK5 SV1) \<or>
@@ -660,7 +660,7 @@
(bnd_s (PUZ114_5_bnd_sK5 SV1))))) =
True) \<or>
SV1 PUZ114_5_bnd_sK1 PUZ114_5_bnd_sK2 = True"
-by (tactic {*full_extcnf_combined_tac @{context}*})
+by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
lemma "\<forall>SV2. (\<forall>SY41.
\<not> PUZ114_5_bnd_sK3 SV2 SY41 \<or>
@@ -670,7 +670,7 @@
(\<not> PUZ114_5_bnd_sK3 SV2 SV4 \<or>
PUZ114_5_bnd_sK3 (dud_bnd_s (dud_bnd_s SV2)) SV4) =
True"
-by (tactic {*full_extcnf_combined_tac @{context}*})
+by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
lemma "\<forall>SV3. (\<forall>SY42.
\<not> PUZ114_5_bnd_sK3 SV3 SY42 \<or>
@@ -680,7 +680,7 @@
(\<not> PUZ114_5_bnd_sK3 SV3 SV5 \<or>
PUZ114_5_bnd_sK3 (dud_bnd_s SV3) (dud_bnd_s SV5)) =
True"
-by (tactic {*full_extcnf_combined_tac @{context}*})
+by (tactic \<open>full_extcnf_combined_tac @{context}\<close>)
subsection "unfold_def"
@@ -749,7 +749,7 @@
(bnd_setadjoin (bnd_setadjoin SX2 (bnd_setadjoin SX3 bnd_emptyset))
bnd_emptyset))))))) =
True"
-by (tactic {*unfold_def_tac @{context} []*})
+by (tactic \<open>unfold_def_tac @{context} []\<close>)
(* (Annotated_step ("10", "unfold_def"), *)
lemma "bnd_kpairiskpair =
@@ -817,7 +817,7 @@
(bnd_setadjoin (bnd_setadjoin SX2 (bnd_setadjoin SX3 bnd_emptyset))
bnd_emptyset)))))) =
True"
-by (tactic {*unfold_def_tac @{context} []*})
+by (tactic \<open>unfold_def_tac @{context} []\<close>)
(* (Annotated_step ("12", "unfold_def"), *)
lemma "bnd_cCKB6_BLACK =
@@ -854,7 +854,7 @@
apply (tactic {*log_expander 1*})+
apply (rule refl)
*)
-by (tactic {*unfold_def_tac @{context} []*})
+by (tactic \<open>unfold_def_tac @{context} []\<close>)
(* (Annotated_step ("13", "unfold_def"), *)
lemma "bnd_cCKB6_BLACK =
@@ -887,10 +887,10 @@
apply (tactic {*expander_animal 1*})+
apply assumption
*)
-by (tactic {*unfold_def_tac @{context} []*})
+by (tactic \<open>unfold_def_tac @{context} []\<close>)
(*FIXME move this heuristic elsewhere*)
-ML {*
+ML \<open>
(*Other than the list (which must not be empty) this function
expects a parameter indicating the smallest integer.
(Using Int.minInt isn't always viable).*)
@@ -905,13 +905,13 @@
let
val max = max_int_floored min l
in find_index (pair max #> op =) l end
-*}
+\<close>
-ML {*
+ML \<open>
max_index_floored 0 [1, 3, 5]
-*}
+\<close>
-ML {*
+ML \<open>
(*
Given argument ([h_1, ..., h_n], conc),
obtained from term of form
@@ -925,9 +925,9 @@
map size_of_term hypos
|> max_index_floored 0
|> SOME
-*}
+\<close>
-ML {*
+ML \<open>
fun biggest_hyp_first_tac i = fn st =>
let
val results = TERMFUN biggest_hypothesis (SOME i) st
@@ -944,7 +944,7 @@
if n > 0 then rotate_tac n i st else no_tac st
end
end
-*}
+\<close>
(* (Annotated_step ("6", "unfold_def"), *)
lemma "(\<not> (\<exists>U :: TPTP_Interpret.ind \<Rightarrow> bool. \<forall>V. U V = SEV405_5_bnd_cA)) = True \<Longrightarrow>
@@ -1006,7 +1006,7 @@
GEG007_1_bnd_c (GEG007_1_bnd_sK7_SX2 SV13 SV6) GEG007_1_bnd_catalunya) =
False \<or>
GEG007_1_bnd_a SV6 SV13 = False"
-by (tactic {*nonfull_extcnf_combined_tac @{context} [Existential_Var]*})
+by (tactic \<open>nonfull_extcnf_combined_tac @{context} [Existential_Var]\<close>)
(* (Annotated_step ("116", "extcnf_forall_neg"), *)
lemma "\<forall>SV13 SV6.
@@ -1039,7 +1039,7 @@
\<not> (\<forall>SX4. \<not> GEG007_1_bnd_c SX4 SY71 \<or> GEG007_1_bnd_c SX4 GEG007_1_bnd_spain)))))) =
False \<or>
GEG007_1_bnd_a SV6 SV13 = False"
-by (tactic {*nonfull_extcnf_combined_tac @{context} [Existential_Var]*})
+by (tactic \<open>nonfull_extcnf_combined_tac @{context} [Existential_Var]\<close>)
consts PUZ107_5_bnd_sK1_SX0 ::
"TPTP_Interpret.ind
@@ -1056,7 +1056,7 @@
(SV3::TPTP_Interpret.ind) SV1::TPTP_Interpret.ind.
((SV1 = SV3) = True \<or> PUZ107_5_bnd_sK1_SX0 SV1 SV2 SV6 SV8 = False) \<or>
PUZ107_5_bnd_sK1_SX0 SV3 SV4 SV6 SV8 = False"
-by (tactic {*nonfull_extcnf_combined_tac @{context} [Not_neg]*})
+by (tactic \<open>nonfull_extcnf_combined_tac @{context} [Not_neg]\<close>)
lemma "
\<forall>(SV8::TPTP_Interpret.ind) (SV6::TPTP_Interpret.ind)
@@ -1070,7 +1070,7 @@
(SV3::TPTP_Interpret.ind) SV1::TPTP_Interpret.ind.
((SV1 \<noteq> SV3) = False \<or> PUZ107_5_bnd_sK1_SX0 SV1 SV2 SV6 SV8 = False) \<or>
PUZ107_5_bnd_sK1_SX0 SV3 SV4 SV6 SV8 = False"
-by (tactic {*nonfull_extcnf_combined_tac @{context} [Or_neg]*})
+by (tactic \<open>nonfull_extcnf_combined_tac @{context} [Or_neg]\<close>)
consts
NUM016_5_bnd_a :: TPTP_Interpret.ind
@@ -1149,7 +1149,7 @@
NUM016_5_bnd_less (NUM016_5_bnd_factorial_plus_one NUM016_5_bnd_a)
SX0))) =
True"
-by (tactic {*unfold_def_tac @{context} []*})
+by (tactic \<open>unfold_def_tac @{context} []\<close>)
(* (Annotated_step ("3", "unfold_def"), *)
lemma "(~ ((((((((((((ALL X. ~ bnd_less X X) &
@@ -1189,13 +1189,13 @@
(ALL X. (~ bnd_prime X | ~ bnd_less bnd_a X) |
bnd_less (bnd_factorial_plus_one bnd_a) X))) =
False"
-by (tactic {*unfold_def_tac @{context} []*})
+by (tactic \<open>unfold_def_tac @{context} []\<close>)
(* SET062^6.p.out
[[(Annotated_step ("3", "unfold_def"), *)
lemma "(\<forall>Z3. False \<longrightarrow> bnd_cA Z3) = False \<Longrightarrow>
(\<forall>Z3. False \<longrightarrow> bnd_cA Z3) = False"
-by (tactic {*unfold_def_tac @{context} []*})
+by (tactic \<open>unfold_def_tac @{context} []\<close>)
(*
(* SEU559^2.p.out *)
@@ -1253,7 +1253,7 @@
False \<Longrightarrow>
AGT037_2_bnd_possibly_likes AGT037_2_bnd_jan AGT037_2_bnd_cola AGT037_2_bnd_sK1_SX0 =
False"
-by (tactic {*nonfull_extcnf_combined_tac @{context} [Existential_Var]*})
+by (tactic \<open>nonfull_extcnf_combined_tac @{context} [Existential_Var]\<close>)
(* (Annotated_step ("202", "extcnf_forall_neg"), *)
lemma "\<forall>(SV13::TPTP_Interpret.ind) (SV39::AGT037_2_bnd_mu) (SV29::AGT037_2_bnd_mu)
@@ -1297,7 +1297,7 @@
apply (tactic {*clause_breaker 1*})
apply (tactic {*nonfull_extcnf_combined_tac []*})
*)
-by (tactic {*nonfull_extcnf_combined_tac @{context} [Existential_Var]*})
+by (tactic \<open>nonfull_extcnf_combined_tac @{context} [Existential_Var]\<close>)
(*(*NUM667^1*)
lemma "\<forall>SV12 SV13 SV14 SV9 SV10 SV11.
@@ -2157,7 +2157,7 @@
\<not> (\<not> SEV405_5_bnd_cA \<or> SV1 (SEV405_5_bnd_sK1_SY2 SV1)))) =
False"
(* apply (tactic {*full_extcnf_combined_tac*}) *)
-by (tactic {*nonfull_extcnf_combined_tac @{context} [Existential_Var]*})
+by (tactic \<open>nonfull_extcnf_combined_tac @{context} [Existential_Var]\<close>)
(* (Annotated_step ("13", "extcnf_forall_pos"), *)
lemma "(\<forall>SY31 SY32.
@@ -2168,7 +2168,7 @@
bnd_sK2 (bnd_sK4 SV1 SY42) =
bnd_sK5 (bnd_sK2 SV1) (bnd_sK2 SY42)) =
True"
-by (tactic {*nonfull_extcnf_combined_tac @{context} [Universal]*})
+by (tactic \<open>nonfull_extcnf_combined_tac @{context} [Universal]\<close>)
(* (Annotated_step ("14", "extcnf_forall_pos"), *)
lemma "(\<forall>SY35 SY36.
@@ -2179,7 +2179,7 @@
bnd_sK1 (bnd_sK3 SV2 SY43) =
bnd_sK4 (bnd_sK1 SV2) (bnd_sK1 SY43)) =
True"
-by (tactic {*nonfull_extcnf_combined_tac @{context} [Universal]*})
+by (tactic \<open>nonfull_extcnf_combined_tac @{context} [Universal]\<close>)
(*from SYO198^5.p.out*)
@@ -2188,7 +2188,7 @@
\<not> \<not> (\<not> SX0 bnd_sK1_Xx \<or> \<not> SX0 bnd_sK2_Xy)) =
True \<Longrightarrow>
(\<not> \<not> (\<not> True \<or> \<not> True)) = True"
-apply (tactic {*extcnf_forall_special_pos_tac 1*})
+apply (tactic \<open>extcnf_forall_special_pos_tac 1\<close>)
done
(* (Annotated_step ("13", "extcnf_forall_special_pos"), *)
@@ -2196,7 +2196,7 @@
\<not> \<not> (\<not> SX0 bnd_sK1_Xx \<or> \<not> SX0 bnd_sK2_Xy)) =
True \<Longrightarrow>
(\<not> \<not> (\<not> bnd_sK1_Xx \<or> \<not> bnd_sK2_Xy)) = True"
-apply (tactic {*extcnf_forall_special_pos_tac 1*})
+apply (tactic \<open>extcnf_forall_special_pos_tac 1\<close>)
done
(* [[(Annotated_step ("8", "polarity_switch"), *)
@@ -2205,7 +2205,7 @@
(\<not> (\<forall>(Xx::bool) (Xy::bool) Xz::bool.
True \<and> True \<longrightarrow> True)) =
True"
-apply (tactic {*nonfull_extcnf_combined_tac @{context} [Polarity_switch]*})
+apply (tactic \<open>nonfull_extcnf_combined_tac @{context} [Polarity_switch]\<close>)
done
lemma "((\<forall>SY22 SY23 SY24.
@@ -2224,7 +2224,7 @@
(\<forall>SY28. SY25 SY28 \<longrightarrow> bnd_sK1_RRR SY28 SY27) \<longrightarrow>
bnd_sK1_RRR (bnd_sK2_U SY25) SY27)) =
True"
-apply (tactic {*standard_cnf_tac @{context} 1*})
+apply (tactic \<open>standard_cnf_tac @{context} 1\<close>)
done
lemma "((\<forall>Xx. bnd_in Xx bnd_emptyset \<longrightarrow> (\<forall>Xphi. Xphi)) \<longrightarrow>
@@ -2236,7 +2236,7 @@
False \<Longrightarrow>
(\<forall>Xx. bnd_in Xx bnd_emptyset \<longrightarrow> (\<forall>Xphi. Xphi)) =
True"
-apply (tactic {*standard_cnf_tac @{context} 1*})
+apply (tactic \<open>standard_cnf_tac @{context} 1\<close>)
done
lemma "((\<forall>Xx. bnd_in Xx bnd_emptyset \<longrightarrow> (\<forall>Xphi. Xphi)) \<longrightarrow>
@@ -2248,7 +2248,7 @@
False \<Longrightarrow>
(\<forall>Xx Xy. bnd_in Xx (bnd_setadjoin Xx Xy)) =
True"
-apply (tactic {*standard_cnf_tac @{context} 1*})
+apply (tactic \<open>standard_cnf_tac @{context} 1\<close>)
done
lemma "((\<forall>Xx. bnd_in Xx bnd_emptyset \<longrightarrow> (\<forall>Xphi. Xphi)) \<longrightarrow>
@@ -2261,7 +2261,7 @@
(\<forall>A B. A = B \<longrightarrow>
(\<forall>Xx Xy. Xx = Xy \<longrightarrow> bnd_in Xx A = bnd_in Xy B)) =
True"
-apply (tactic {*standard_cnf_tac @{context} 1*})
+apply (tactic \<open>standard_cnf_tac @{context} 1\<close>)
done
lemma "((\<forall>Xx. bnd_in Xx bnd_emptyset \<longrightarrow> (\<forall>Xphi. Xphi)) \<longrightarrow>
@@ -2274,7 +2274,7 @@
(\<forall>SY0. bnd_in SY0 bnd_omega \<longrightarrow>
bnd_setadjoin SY0 SY0 \<noteq> bnd_emptyset) =
False"
-apply (tactic {*standard_cnf_tac @{context} 1*})
+apply (tactic \<open>standard_cnf_tac @{context} 1\<close>)
done
(*nested conjunctions*)
@@ -2287,7 +2287,7 @@
False \<Longrightarrow>
(\<forall>Xx. bnd_cP bnd_e Xx = Xx) =
True"
-apply (tactic {*standard_cnf_tac @{context} 1*})
+apply (tactic \<open>standard_cnf_tac @{context} 1\<close>)
done
end
--- a/src/HOL/TPTP/TPTP_Test.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/TPTP/TPTP_Test.thy Thu May 26 17:51:22 2016 +0200
@@ -9,7 +9,7 @@
imports TPTP_Parser
begin
-ML {*
+ML \<open>
val tptp_test_out = Attrib.setup_config_string @{binding "tptp_test_out"} (K "")
val tptp_test_all = Attrib.setup_config_bool @{binding "tptp_test_all"} (K false)
val tptp_test_timeout =
@@ -17,27 +17,27 @@
fun test_all ctxt = Config.get ctxt tptp_test_all
fun get_timeout ctxt = Config.get ctxt tptp_test_timeout
fun S x y z = x z (y z)
-*}
+\<close>
section "Parser tests"
-ML {*
+ML \<open>
fun payload_of (TPTP_Syntax.Annotated_Formula (_, _, _, _, fmla, _)) = fmla
val payloads_of = map payload_of
-*}
+\<close>
section "Source problems"
-ML {*
+ML \<open>
(*problem source*)
val tptp_probs_dir =
Path.explode "$TPTP/Problems"
|> Path.expand;
-*}
+\<close>
section "Supporting test functions"
-ML {*
+ML \<open>
fun report ctxt str =
let
val tptp_test_out = Config.get ctxt tptp_test_out
@@ -85,10 +85,10 @@
report ctxt ("Test timing: " ^ Timing.message time ^ "\n(about " ^ average ^
"s per problem)")
end
-*}
+\<close>
-ML {*
+ML \<open>
fun situate file_name = Path.append tptp_probs_dir (Path.explode file_name);
fun parser_test ctxt = (*FIXME argument order*)
@@ -103,6 +103,6 @@
fun parse_timed file =
Timing.timing TPTP_Parser.parse_file (Path.implode file)
-*}
+\<close>
end
--- a/src/HOL/ZF/HOLZF.thy Thu May 26 16:57:14 2016 +0200
+++ b/src/HOL/ZF/HOLZF.thy Thu May 26 17:51:22 2016 +0200
@@ -188,7 +188,7 @@
definition Field :: "ZF \<Rightarrow> ZF" where
"Field A == union (Domain A) (Range A)"
-definition app :: "ZF \<Rightarrow> ZF => ZF" (infixl "\<acute>" 90) --{*function application*} where
+definition app :: "ZF \<Rightarrow> ZF => ZF" (infixl "\<acute>" 90) \<comment>\<open>function application\<close> where
"f \<acute> x == (THE y. Elem (Opair x y) f)"
definition isFun :: "ZF \<Rightarrow> bool" where