--- a/src/CCL/CCL.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/CCL/CCL.thy Thu Jul 23 14:25:05 2015 +0200
@@ -3,18 +3,18 @@
Copyright 1993 University of Cambridge
*)
-section {* Classical Computational Logic for Untyped Lambda Calculus
- with reduction to weak head-normal form *}
+section \<open>Classical Computational Logic for Untyped Lambda Calculus
+ with reduction to weak head-normal form\<close>
theory CCL
imports Gfp
begin
-text {*
+text \<open>
Based on FOL extended with set collection, a primitive higher-order
logic. HOL is too strong - descriptions prevent a type of programs
being defined which contains only executable terms.
-*}
+\<close>
class prog = "term"
default_sort prog
@@ -150,7 +150,7 @@
definition Trm :: "i \<Rightarrow> o"
where "Trm(t) == \<not> Dvg(t)"
-text {*
+text \<open>
Would be interesting to build a similar theory for a typed programming language:
ie. true :: bool, fix :: ('a\<Rightarrow>'a)\<Rightarrow>'a etc......
@@ -158,7 +158,7 @@
What are the advantages of this approach?
- less axiomatic
- wfd induction / coinduction and fixed point induction available
-*}
+\<close>
lemmas ccl_data_defs = apply_def fix_def
@@ -166,7 +166,7 @@
declare po_refl [simp]
-subsection {* Congruence Rules *}
+subsection \<open>Congruence Rules\<close>
(*similar to AP_THM in Gordon's HOL*)
lemma fun_cong: "(f::'a\<Rightarrow>'b) = g \<Longrightarrow> f(x)=g(x)"
@@ -184,7 +184,7 @@
lemmas caseBs = caseBtrue caseBfalse caseBpair caseBlam caseBbot
-subsection {* Termination and Divergence *}
+subsection \<open>Termination and Divergence\<close>
lemma Trm_iff: "Trm(t) \<longleftrightarrow> \<not> t = bot"
by (simp add: Trm_def Dvg_def)
@@ -193,12 +193,12 @@
by (simp add: Trm_def Dvg_def)
-subsection {* Constructors are injective *}
+subsection \<open>Constructors are injective\<close>
lemma eq_lemma: "\<lbrakk>x=a; y=b; x=y\<rbrakk> \<Longrightarrow> a=b"
by simp
-ML {*
+ML \<open>
fun inj_rl_tac ctxt rews i =
let
fun mk_inj_lemmas r = [@{thm arg_cong}] RL [r RS (r RS @{thm eq_lemma})]
@@ -209,11 +209,11 @@
eresolve_tac ctxt inj_lemmas i ORELSE
asm_simp_tac (ctxt addsimps rews) i))
end;
-*}
+\<close>
-method_setup inj_rl = {*
+method_setup inj_rl = \<open>
Attrib.thms >> (fn rews => fn ctxt => SIMPLE_METHOD' (inj_rl_tac ctxt rews))
-*}
+\<close>
lemma ccl_injs:
"<a,b> = <a',b'> \<longleftrightarrow> (a=a' \<and> b=b')"
@@ -225,12 +225,12 @@
by (simp add: ccl_injs)
-subsection {* Constructors are distinct *}
+subsection \<open>Constructors are distinct\<close>
lemma lem: "t=t' \<Longrightarrow> case(t,b,c,d,e) = case(t',b,c,d,e)"
by simp
-ML {*
+ML \<open>
local
fun pairs_of f x [] = []
| pairs_of f x (y::ys) = (f x y) :: (f y x) :: (pairs_of f x ys)
@@ -259,9 +259,9 @@
fun mk_lemmas rls = maps mk_lemma (mk_combs pair rls)
fun mk_dstnct_rls thy xs = mk_combs (mk_thm_str thy) xs
end
-*}
+\<close>
-ML {*
+ML \<open>
val caseB_lemmas = mk_lemmas @{thms caseBs}
val ccl_dstncts =
@@ -296,14 +296,14 @@
ML_Thms.bind_thms ("ccl_rews", @{thms caseBs} @ @{thms ccl_injs} @ ccl_dstncts);
ML_Thms.bind_thms ("ccl_dstnctsEs", ccl_dstncts RL [@{thm notE}]);
ML_Thms.bind_thms ("ccl_injDs", XH_to_Ds @{thms ccl_injs});
-*}
+\<close>
lemmas [simp] = ccl_rews
and [elim!] = pair_inject ccl_dstnctsEs
and [dest!] = ccl_injDs
-subsection {* Facts from gfp Definition of @{text "[="} and @{text "="} *}
+subsection \<open>Facts from gfp Definition of @{text "[="} and @{text "="}\<close>
lemma XHlemma1: "\<lbrakk>A=B; a:B \<longleftrightarrow> P\<rbrakk> \<Longrightarrow> a:A \<longleftrightarrow> P"
by simp
@@ -312,7 +312,7 @@
by blast
-subsection {* Pre-Order *}
+subsection \<open>Pre-Order\<close>
lemma POgen_mono: "mono(\<lambda>X. POgen(X))"
apply (unfold POgen_def SIM_def)
@@ -422,7 +422,7 @@
lemmas npo_rls = npo_pair_lam npo_lam_pair npo_rls1
-subsection {* Coinduction for @{text "[="} *}
+subsection \<open>Coinduction for @{text "[="}\<close>
lemma po_coinduct: "\<lbrakk><t,u> : R; R <= POgen(R)\<rbrakk> \<Longrightarrow> t [= u"
apply (rule PO_def [THEN def_coinduct, THEN PO_iff [THEN iffD2]])
@@ -430,7 +430,7 @@
done
-subsection {* Equality *}
+subsection \<open>Equality\<close>
lemma EQgen_mono: "mono(\<lambda>X. EQgen(X))"
apply (unfold EQgen_def SIM_def)
@@ -473,14 +473,14 @@
apply (rule EQgen_mono | assumption)+
done
-method_setup eq_coinduct3 = {*
+method_setup eq_coinduct3 = \<open>
Scan.lift Args.name_inner_syntax >> (fn s => fn ctxt =>
SIMPLE_METHOD'
(Rule_Insts.res_inst_tac ctxt [((("R", 0), Position.none), s)] [] @{thm eq_coinduct3}))
-*}
+\<close>
-subsection {* Untyped Case Analysis and Other Facts *}
+subsection \<open>Untyped Case Analysis and Other Facts\<close>
lemma cond_eta: "(EX f. t=lam x. f(x)) \<Longrightarrow> t = lam x.(t ` x)"
by (auto simp: apply_def)
--- a/src/CCL/Fix.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/CCL/Fix.thy Thu Jul 23 14:25:05 2015 +0200
@@ -3,7 +3,7 @@
Copyright 1993 University of Cambridge
*)
-section {* Tentative attempt at including fixed point induction; justified by Smith *}
+section \<open>Tentative attempt at including fixed point induction; justified by Smith\<close>
theory Fix
imports Type
@@ -18,7 +18,7 @@
INCL_subst: "INCL(P) \<Longrightarrow> INCL(\<lambda>x. P((g::i\<Rightarrow>i)(x)))"
-subsection {* Fixed Point Induction *}
+subsection \<open>Fixed Point Induction\<close>
lemma fix_ind:
assumes base: "P(bot)"
@@ -33,7 +33,7 @@
done
-subsection {* Inclusive Predicates *}
+subsection \<open>Inclusive Predicates\<close>
lemma inclXH: "INCL(P) \<longleftrightarrow> (ALL f. (ALL n:Nat. P(f ^ n ` bot)) \<longrightarrow> P(fix(f)))"
by (simp add: INCL_def)
@@ -48,7 +48,7 @@
by (blast dest: inclD)
-subsection {* Lemmas for Inclusive Predicates *}
+subsection \<open>Lemmas for Inclusive Predicates\<close>
lemma npo_INCL: "INCL(\<lambda>x. \<not> a(x) [= t)"
apply (rule inclI)
@@ -77,7 +77,7 @@
done
-subsection {* Derivation of Reachability Condition *}
+subsection \<open>Derivation of Reachability Condition\<close>
(* Fixed points of idgen *)
--- a/src/CCL/Gfp.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/CCL/Gfp.thy Thu Jul 23 14:25:05 2015 +0200
@@ -3,7 +3,7 @@
Copyright 1992 University of Cambridge
*)
-section {* Greatest fixed points *}
+section \<open>Greatest fixed points\<close>
theory Gfp
imports Lfp
@@ -90,7 +90,7 @@
done
-subsection {* Definition forms of @{text "gfp_Tarski"}, to control unfolding *}
+subsection \<open>Definition forms of @{text "gfp_Tarski"}, to control unfolding\<close>
lemma def_gfp_Tarski: "\<lbrakk>h == gfp(f); mono(f)\<rbrakk> \<Longrightarrow> h = f(h)"
apply unfold
--- a/src/CCL/Hered.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/CCL/Hered.thy Thu Jul 23 14:25:05 2015 +0200
@@ -3,17 +3,17 @@
Copyright 1993 University of Cambridge
*)
-section {* Hereditary Termination -- cf. Martin Lo\"f *}
+section \<open>Hereditary Termination -- cf. Martin Lo\"f\<close>
theory Hered
imports Type
begin
-text {*
+text \<open>
Note that this is based on an untyped equality and so @{text "lam
x. b(x)"} is only hereditarily terminating if @{text "ALL x. b(x)"}
is. Not so useful for functions!
-*}
+\<close>
definition HTTgen :: "i set \<Rightarrow> i set" where
"HTTgen(R) ==
@@ -24,7 +24,7 @@
where "HTT == gfp(HTTgen)"
-subsection {* Hereditary Termination *}
+subsection \<open>Hereditary Termination\<close>
lemma HTTgen_mono: "mono(\<lambda>X. HTTgen(X))"
apply (unfold HTTgen_def)
@@ -47,7 +47,7 @@
done
-subsection {* Introduction Rules for HTT *}
+subsection \<open>Introduction Rules for HTT\<close>
lemma HTT_bot: "\<not> bot : HTT"
by (blast dest: HTTXH [THEN iffD1])
@@ -83,7 +83,7 @@
lemmas HTT_rews = HTT_rews1 HTT_rews2
-subsection {* Coinduction for HTT *}
+subsection \<open>Coinduction for HTT\<close>
lemma HTT_coinduct: "\<lbrakk>t : R; R <= HTTgen(R)\<rbrakk> \<Longrightarrow> t : HTT"
apply (erule HTT_def [THEN def_coinduct])
@@ -111,7 +111,7 @@
unfolding data_defs by (genIs HTTgenXH HTTgen_mono)+
-subsection {* Formation Rules for Types *}
+subsection \<open>Formation Rules for Types\<close>
lemma UnitF: "Unit <= HTT"
by (simp add: subsetXH UnitXH HTT_rews)
--- a/src/CCL/Lfp.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/CCL/Lfp.thy Thu Jul 23 14:25:05 2015 +0200
@@ -3,7 +3,7 @@
Copyright 1992 University of Cambridge
*)
-section {* The Knaster-Tarski Theorem *}
+section \<open>The Knaster-Tarski Theorem\<close>
theory Lfp
imports Set
--- a/src/CCL/Set.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/CCL/Set.thy Thu Jul 23 14:25:05 2015 +0200
@@ -1,4 +1,4 @@
-section {* Extending FOL by a modified version of HOL set theory *}
+section \<open>Extending FOL by a modified version of HOL set theory\<close>
theory Set
imports "~~/src/FOL/FOL"
@@ -83,7 +83,7 @@
done
-subsection {* Bounded quantifiers *}
+subsection \<open>Bounded quantifiers\<close>
lemma ballI: "(\<And>x. x:A \<Longrightarrow> P(x)) \<Longrightarrow> ALL x:A. P(x)"
by (simp add: Ball_def)
@@ -108,7 +108,7 @@
by (blast intro: ballI)
-subsection {* Congruence rules *}
+subsection \<open>Congruence rules\<close>
lemma ball_cong:
"\<lbrakk>A = A'; \<And>x. x:A' \<Longrightarrow> P(x) \<longleftrightarrow> P'(x)\<rbrakk> \<Longrightarrow>
@@ -121,7 +121,7 @@
by (blast intro: bexI elim: bexE)
-subsection {* Rules for subsets *}
+subsection \<open>Rules for subsets\<close>
lemma subsetI: "(\<And>x. x:A \<Longrightarrow> x:B) \<Longrightarrow> A <= B"
unfolding subset_def by (blast intro: ballI)
@@ -141,7 +141,7 @@
by (blast intro: subsetI dest: subsetD)
-subsection {* Rules for equality *}
+subsection \<open>Rules for equality\<close>
(*Anti-symmetry of the subset relation*)
lemma subset_antisym: "\<lbrakk>A <= B; B <= A\<rbrakk> \<Longrightarrow> A = B"
@@ -164,7 +164,7 @@
by (blast intro: equalityI subsetI CollectI dest: CollectD)
-subsection {* Rules for binary union *}
+subsection \<open>Rules for binary union\<close>
lemma UnI1: "c:A \<Longrightarrow> c : A Un B"
and UnI2: "c:B \<Longrightarrow> c : A Un B"
@@ -178,7 +178,7 @@
unfolding Un_def by (blast dest: CollectD)
-subsection {* Rules for small intersection *}
+subsection \<open>Rules for small intersection\<close>
lemma IntI: "\<lbrakk>c:A; c:B\<rbrakk> \<Longrightarrow> c : A Int B"
unfolding Int_def by (blast intro: CollectI)
@@ -191,7 +191,7 @@
by (blast dest: IntD1 IntD2)
-subsection {* Rules for set complement *}
+subsection \<open>Rules for set complement\<close>
lemma ComplI: "(c:A \<Longrightarrow> False) \<Longrightarrow> c : Compl(A)"
unfolding Compl_def by (blast intro: CollectI)
@@ -205,7 +205,7 @@
lemmas ComplE = ComplD [elim_format]
-subsection {* Empty sets *}
+subsection \<open>Empty sets\<close>
lemma empty_eq: "{x. False} = {}"
by (simp add: empty_def)
@@ -225,7 +225,7 @@
qed
-subsection {* Singleton sets *}
+subsection \<open>Singleton sets\<close>
lemma singletonI: "a : {a}"
unfolding singleton_def by (blast intro: CollectI)
@@ -236,7 +236,7 @@
lemmas singletonE = singletonD [elim_format]
-subsection {* Unions of families *}
+subsection \<open>Unions of families\<close>
(*The order of the premises presupposes that A is rigid; b may be flexible*)
lemma UN_I: "\<lbrakk>a:A; b: B(a)\<rbrakk> \<Longrightarrow> b: (UN x:A. B(x))"
@@ -249,7 +249,7 @@
by (simp add: UNION_def cong: bex_cong)
-subsection {* Intersections of families *}
+subsection \<open>Intersections of families\<close>
lemma INT_I: "(\<And>x. x:A \<Longrightarrow> b: B(x)) \<Longrightarrow> b : (INT x:A. B(x))"
unfolding INTER_def by (blast intro: CollectI ballI)
@@ -265,7 +265,7 @@
by (simp add: INTER_def cong: ball_cong)
-subsection {* Rules for Unions *}
+subsection \<open>Rules for Unions\<close>
(*The order of the premises presupposes that C is rigid; A may be flexible*)
lemma UnionI: "\<lbrakk>X:C; A:X\<rbrakk> \<Longrightarrow> A : Union(C)"
@@ -275,7 +275,7 @@
unfolding Union_def by (blast elim: UN_E)
-subsection {* Rules for Inter *}
+subsection \<open>Rules for Inter\<close>
lemma InterI: "(\<And>X. X:C \<Longrightarrow> A:X) \<Longrightarrow> A : Inter(C)"
unfolding Inter_def by (blast intro: INT_I)
@@ -290,9 +290,9 @@
unfolding Inter_def by (blast elim: INT_E)
-section {* Derived rules involving subsets; Union and Intersection as lattice operations *}
+section \<open>Derived rules involving subsets; Union and Intersection as lattice operations\<close>
-subsection {* Big Union -- least upper bound of a set *}
+subsection \<open>Big Union -- least upper bound of a set\<close>
lemma Union_upper: "B:A \<Longrightarrow> B <= Union(A)"
by (blast intro: subsetI UnionI)
@@ -301,7 +301,7 @@
by (blast intro: subsetI dest: subsetD elim: UnionE)
-subsection {* Big Intersection -- greatest lower bound of a set *}
+subsection \<open>Big Intersection -- greatest lower bound of a set\<close>
lemma Inter_lower: "B:A \<Longrightarrow> Inter(A) <= B"
by (blast intro: subsetI dest: InterD)
@@ -310,7 +310,7 @@
by (blast intro: subsetI InterI dest: subsetD)
-subsection {* Finite Union -- the least upper bound of 2 sets *}
+subsection \<open>Finite Union -- the least upper bound of 2 sets\<close>
lemma Un_upper1: "A <= A Un B"
by (blast intro: subsetI UnI1)
@@ -322,7 +322,7 @@
by (blast intro: subsetI elim: UnE dest: subsetD)
-subsection {* Finite Intersection -- the greatest lower bound of 2 sets *}
+subsection \<open>Finite Intersection -- the greatest lower bound of 2 sets\<close>
lemma Int_lower1: "A Int B <= A"
by (blast intro: subsetI elim: IntE)
@@ -334,7 +334,7 @@
by (blast intro: subsetI IntI dest: subsetD)
-subsection {* Monotonicity *}
+subsection \<open>Monotonicity\<close>
lemma monoI: "(\<And>A B. A <= B \<Longrightarrow> f(A) <= f(B)) \<Longrightarrow> mono(f)"
unfolding mono_def by blast
@@ -349,7 +349,7 @@
by (blast intro: Int_greatest dest: monoD intro: Int_lower1 Int_lower2)
-subsection {* Automated reasoning setup *}
+subsection \<open>Automated reasoning setup\<close>
lemmas [intro!] = ballI subsetI InterI INT_I CollectI ComplI IntI UnCI singletonI
and [intro] = bexI UnionI UN_I
@@ -369,9 +369,9 @@
and [cong] = ball_cong bex_cong INT_cong UN_cong
-section {* Equalities involving union, intersection, inclusion, etc. *}
+section \<open>Equalities involving union, intersection, inclusion, etc.\<close>
-subsection {* Binary Intersection *}
+subsection \<open>Binary Intersection\<close>
lemma Int_absorb: "A Int A = A"
by (blast intro: equalityI)
@@ -389,7 +389,7 @@
by (blast intro: equalityI elim: equalityE)
-subsection {* Binary Union *}
+subsection \<open>Binary Union\<close>
lemma Un_absorb: "A Un A = A"
by (blast intro: equalityI)
@@ -411,7 +411,7 @@
by (blast intro: equalityI elim: equalityE)
-subsection {* Simple properties of @{text "Compl"} -- complement of a set *}
+subsection \<open>Simple properties of @{text "Compl"} -- complement of a set\<close>
lemma Compl_disjoint: "A Int Compl(A) = {x. False}"
by (blast intro: equalityI)
@@ -439,7 +439,7 @@
by (blast intro: equalityI elim: equalityE)
-subsection {* Big Union and Intersection *}
+subsection \<open>Big Union and Intersection\<close>
lemma Union_Un_distrib: "Union(A Un B) = Union(A) Un Union(B)"
by (blast intro: equalityI)
@@ -452,7 +452,7 @@
by (blast intro: equalityI)
-subsection {* Unions and Intersections of Families *}
+subsection \<open>Unions and Intersections of Families\<close>
lemma UN_eq: "(UN x:A. B(x)) = Union({Y. EX x:A. Y=B(x)})"
by (blast intro: equalityI)
@@ -468,7 +468,7 @@
by (blast intro: equalityI)
-section {* Monotonicity of various operations *}
+section \<open>Monotonicity of various operations\<close>
lemma Union_mono: "A<=B \<Longrightarrow> Union(A) <= Union(B)"
by blast
--- a/src/CCL/Term.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/CCL/Term.thy Thu Jul 23 14:25:05 2015 +0200
@@ -3,7 +3,7 @@
Copyright 1993 University of Cambridge
*)
-section {* Definitions of usual program constructs in CCL *}
+section \<open>Definitions of usual program constructs in CCL\<close>
theory Term
imports CCL
@@ -52,7 +52,7 @@
"_letrec3" :: "[id,id,id,id,i,i]\<Rightarrow>i" ("(3letrec _ _ _ _ be _/ in _)"
[0,0,0,0,0,60] 60)
-ML {*
+ML \<open>
(** Quantifier translations: variable binding **)
(* FIXME does not handle "_idtdummy" *)
@@ -91,21 +91,21 @@
val (x',a') = Syntax_Trans.variant_abs(x,T,a3)
in Const(@{syntax_const "_letrec3"},dummyT) $ Free(f',SS) $ Free(x',T) $ Free(y',U) $ Free(z',V) $ a' $ b'
end;
-*}
+\<close>
-parse_translation {*
+parse_translation \<open>
[(@{syntax_const "_let"}, K let_tr),
(@{syntax_const "_letrec"}, K letrec_tr),
(@{syntax_const "_letrec2"}, K letrec2_tr),
(@{syntax_const "_letrec3"}, K letrec3_tr)]
-*}
+\<close>
-print_translation {*
+print_translation \<open>
[(@{const_syntax let}, K let_tr'),
(@{const_syntax letrec}, K letrec_tr'),
(@{const_syntax letrec2}, K letrec2_tr'),
(@{const_syntax letrec3}, K letrec3_tr')]
-*}
+\<close>
consts
napply :: "[i\<Rightarrow>i,i,i]\<Rightarrow>i" ("(_ ^ _ ` _)" [56,56,56] 56)
@@ -156,7 +156,7 @@
and genrec_defs = letrec_def letrec2_def letrec3_def
-subsection {* Beta Rules, including strictness *}
+subsection \<open>Beta Rules, including strictness\<close>
lemma letB: "\<not> t=bot \<Longrightarrow> let x be t in f(x) = f(t)"
apply (unfold let_def)
@@ -200,11 +200,11 @@
lemmas rawBs = caseBs applyB applyBbot
-method_setup beta_rl = {*
+method_setup beta_rl = \<open>
Scan.succeed (fn ctxt =>
SIMPLE_METHOD' (CHANGED o
simp_tac (ctxt addsimps @{thms rawBs} setloop (fn _ => stac ctxt @{thm letrecB}))))
-*}
+\<close>
lemma ifBtrue: "if true then t else u = t"
and ifBfalse: "if false then t else u = u"
@@ -272,7 +272,7 @@
napplyBzero napplyBsucc
-subsection {* Constructors are injective *}
+subsection \<open>Constructors are injective\<close>
lemma term_injs:
"(inl(a) = inl(a')) \<longleftrightarrow> (a=a')"
@@ -282,16 +282,16 @@
by (inj_rl applyB splitB whenBinl whenBinr ncaseBsucc lcaseBcons)
-subsection {* Constructors are distinct *}
+subsection \<open>Constructors are distinct\<close>
-ML {*
+ML \<open>
ML_Thms.bind_thms ("term_dstncts",
mkall_dstnct_thms @{context} @{thms data_defs} (@{thms ccl_injs} @ @{thms term_injs})
[["bot","inl","inr"], ["bot","zero","succ"], ["bot","nil","cons"]]);
-*}
+\<close>
-subsection {* Rules for pre-order @{text "[="} *}
+subsection \<open>Rules for pre-order @{text "[="}\<close>
lemma term_porews:
"inl(a) [= inl(a') \<longleftrightarrow> a [= a'"
@@ -301,11 +301,11 @@
by (simp_all add: data_defs ccl_porews)
-subsection {* Rewriting and Proving *}
+subsection \<open>Rewriting and Proving\<close>
-ML {*
+ML \<open>
ML_Thms.bind_thms ("term_injDs", XH_to_Ds @{thms term_injs});
-*}
+\<close>
lemmas term_rews = termBs term_injs term_dstncts ccl_porews term_porews
--- a/src/CCL/Trancl.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/CCL/Trancl.thy Thu Jul 23 14:25:05 2015 +0200
@@ -3,7 +3,7 @@
Copyright 1993 University of Cambridge
*)
-section {* Transitive closure of a relation *}
+section \<open>Transitive closure of a relation\<close>
theory Trancl
imports CCL
@@ -25,7 +25,7 @@
where "r^+ == r O rtrancl(r)"
-subsection {* Natural deduction for @{text "trans(r)"} *}
+subsection \<open>Natural deduction for @{text "trans(r)"}\<close>
lemma transI: "(\<And>x y z. \<lbrakk><x,y>:r; <y,z>:r\<rbrakk> \<Longrightarrow> <x,z>:r) \<Longrightarrow> trans(r)"
unfolding trans_def by blast
@@ -34,7 +34,7 @@
unfolding trans_def by blast
-subsection {* Identity relation *}
+subsection \<open>Identity relation\<close>
lemma idI: "<a,a> : id"
apply (unfold id_def)
@@ -50,7 +50,7 @@
done
-subsection {* Composition of two relations *}
+subsection \<open>Composition of two relations\<close>
lemma compI: "\<lbrakk><a,b>:s; <b,c>:r\<rbrakk> \<Longrightarrow> <a,c> : r O s"
unfolding relcomp_def by blast
@@ -72,7 +72,7 @@
by blast
-subsection {* The relation rtrancl *}
+subsection \<open>The relation rtrancl\<close>
lemma rtrancl_fun_mono: "mono(\<lambda>s. id Un (r O s))"
apply (rule monoI)
@@ -102,7 +102,7 @@
done
-subsection {* standard induction rule *}
+subsection \<open>standard induction rule\<close>
lemma rtrancl_full_induct:
"\<lbrakk><a,b> : r^*;
@@ -149,9 +149,9 @@
done
-subsection {* The relation trancl *}
+subsection \<open>The relation trancl\<close>
-subsubsection {* Conversions between trancl and rtrancl *}
+subsubsection \<open>Conversions between trancl and rtrancl\<close>
lemma trancl_into_rtrancl: "<a,b> : r^+ \<Longrightarrow> <a,b> : r^*"
apply (unfold trancl_def)
--- a/src/CCL/Type.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/CCL/Type.thy Thu Jul 23 14:25:05 2015 +0200
@@ -3,7 +3,7 @@
Copyright 1993 University of Cambridge
*)
-section {* Types in CCL are defined as sets of terms *}
+section \<open>Types in CCL are defined as sets of terms\<close>
theory Type
imports Term
@@ -45,12 +45,12 @@
"A * B" => "CONST Sigma(A, \<lambda>_. B)"
"{x: A. B}" == "CONST Subtype(A, \<lambda>x. B)"
-print_translation {*
+print_translation \<open>
[(@{const_syntax Pi},
fn _ => Syntax_Trans.dependent_tr' (@{syntax_const "_Pi"}, @{syntax_const "_arrow"})),
(@{const_syntax Sigma},
fn _ => Syntax_Trans.dependent_tr' (@{syntax_const "_Sigma"}, @{syntax_const "_star"}))]
-*}
+\<close>
defs
Subtype_def: "{x:A. P(x)} == {x. x:A \<and> P(x)}"
@@ -82,7 +82,7 @@
by blast
-subsection {* Exhaustion Rules *}
+subsection \<open>Exhaustion Rules\<close>
lemma EmptyXH: "\<And>a. a : {} \<longleftrightarrow> False"
and SubtypeXH: "\<And>a A P. a : {x:A. P(x)} \<longleftrightarrow> (a:A \<and> P(a))"
@@ -100,10 +100,10 @@
and TexXH: "a : TEX X. B(X) \<longleftrightarrow> (EX X. a:B(X))"
unfolding simp_type_defs by blast+
-ML {* ML_Thms.bind_thms ("case_rls", XH_to_Es @{thms XHs}) *}
+ML \<open>ML_Thms.bind_thms ("case_rls", XH_to_Es @{thms XHs})\<close>
-subsection {* Canonical Type Rules *}
+subsection \<open>Canonical Type Rules\<close>
lemma oneT: "one : Unit"
and trueT: "true : Bool"
@@ -117,13 +117,13 @@
lemmas canTs = oneT trueT falseT pairT lamT inlT inrT
-subsection {* Non-Canonical Type Rules *}
+subsection \<open>Non-Canonical Type Rules\<close>
lemma lem: "\<lbrakk>a:B(u); u = v\<rbrakk> \<Longrightarrow> a : B(v)"
by blast
-ML {*
+ML \<open>
fun mk_ncanT_tac top_crls crls =
SUBPROOF (fn {context = ctxt, prems = major :: prems, ...} =>
resolve_tac ctxt ([major] RL top_crls) 1 THEN
@@ -132,11 +132,11 @@
ALLGOALS (assume_tac ctxt ORELSE' resolve_tac ctxt (prems RL [@{thm lem}])
ORELSE' eresolve_tac ctxt @{thms bspec}) THEN
safe_tac (ctxt addSIs prems))
-*}
+\<close>
-method_setup ncanT = {*
+method_setup ncanT = \<open>
Scan.succeed (SIMPLE_METHOD' o mk_ncanT_tac @{thms case_rls} @{thms case_rls})
-*}
+\<close>
lemma ifT: "\<lbrakk>b:Bool; b=true \<Longrightarrow> t:A(true); b=false \<Longrightarrow> u:A(false)\<rbrakk> \<Longrightarrow> if b then t else u : A(b)"
by ncanT
@@ -156,7 +156,7 @@
lemmas ncanTs = ifT applyT splitT whenT
-subsection {* Subtypes *}
+subsection \<open>Subtypes\<close>
lemma SubtypeD1: "a : Subtype(A, P) \<Longrightarrow> a : A"
and SubtypeD2: "a : Subtype(A, P) \<Longrightarrow> P(a)"
@@ -169,7 +169,7 @@
by (simp add: SubtypeXH)
-subsection {* Monotonicity *}
+subsection \<open>Monotonicity\<close>
lemma idM: "mono (\<lambda>X. X)"
apply (rule monoI)
@@ -206,9 +206,9 @@
dest!: monoD [THEN subsetD])
-subsection {* Recursive types *}
+subsection \<open>Recursive types\<close>
-subsubsection {* Conversion Rules for Fixed Points via monotonicity and Tarski *}
+subsubsection \<open>Conversion Rules for Fixed Points via monotonicity and Tarski\<close>
lemma NatM: "mono(\<lambda>X. Unit+X)"
apply (rule PlusM constM idM)+
@@ -245,7 +245,7 @@
lemmas ind_type_eqs = def_NatB def_ListB def_ListsB def_IListsB
-subsection {* Exhaustion Rules *}
+subsection \<open>Exhaustion Rules\<close>
lemma NatXH: "a : Nat \<longleftrightarrow> (a=zero | (EX x:Nat. a=succ(x)))"
and ListXH: "a : List(A) \<longleftrightarrow> (a=[] | (EX x:A. EX xs:List(A).a=x$xs))"
@@ -256,10 +256,10 @@
lemmas iXHs = NatXH ListXH
-ML {* ML_Thms.bind_thms ("icase_rls", XH_to_Es @{thms iXHs}) *}
+ML \<open>ML_Thms.bind_thms ("icase_rls", XH_to_Es @{thms iXHs})\<close>
-subsection {* Type Rules *}
+subsection \<open>Type Rules\<close>
lemma zeroT: "zero : Nat"
and succT: "n:Nat \<Longrightarrow> succ(n) : Nat"
@@ -270,9 +270,9 @@
lemmas icanTs = zeroT succT nilT consT
-method_setup incanT = {*
+method_setup incanT = \<open>
Scan.succeed (SIMPLE_METHOD' o mk_ncanT_tac @{thms icase_rls} @{thms case_rls})
-*}
+\<close>
lemma ncaseT: "\<lbrakk>n:Nat; n=zero \<Longrightarrow> b:C(zero); \<And>x. \<lbrakk>x:Nat; n=succ(x)\<rbrakk> \<Longrightarrow> c(x):C(succ(x))\<rbrakk>
\<Longrightarrow> ncase(n,b,c) : C(n)"
@@ -285,7 +285,7 @@
lemmas incanTs = ncaseT lcaseT
-subsection {* Induction Rules *}
+subsection \<open>Induction Rules\<close>
lemmas ind_Ms = NatM ListM
@@ -304,7 +304,7 @@
lemmas inds = Nat_ind List_ind
-subsection {* Primitive Recursive Rules *}
+subsection \<open>Primitive Recursive Rules\<close>
lemma nrecT: "\<lbrakk>n:Nat; b:C(zero); \<And>x g. \<lbrakk>x:Nat; g:C(x)\<rbrakk> \<Longrightarrow> c(x,g):C(succ(x))\<rbrakk>
\<Longrightarrow> nrec(n,b,c) : C(n)"
@@ -317,7 +317,7 @@
lemmas precTs = nrecT lrecT
-subsection {* Theorem proving *}
+subsection \<open>Theorem proving\<close>
lemma SgE2: "\<lbrakk><a,b> : Sigma(A,B); \<lbrakk>a:A; b:B(a)\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
unfolding SgXH by blast
@@ -326,13 +326,13 @@
(* - intro rules are type rules for canonical terms *)
(* - elim rules are case rules (no non-canonical terms appear) *)
-ML {* ML_Thms.bind_thms ("XHEs", XH_to_Es @{thms XHs}) *}
+ML \<open>ML_Thms.bind_thms ("XHEs", XH_to_Es @{thms XHs})\<close>
lemmas [intro!] = SubtypeI canTs icanTs
and [elim!] = SubtypeE XHEs
-subsection {* Infinite Data Types *}
+subsection \<open>Infinite Data Types\<close>
lemma lfp_subset_gfp: "mono(f) \<Longrightarrow> lfp(f) <= gfp(f)"
apply (rule lfp_lowerbound [THEN subset_trans])
@@ -364,8 +364,8 @@
done
-subsection {* Lemmas and tactics for using the rule @{text
- "coinduct3"} on @{text "[="} and @{text "="} *}
+subsection \<open>Lemmas and tactics for using the rule @{text
+ "coinduct3"} on @{text "[="} and @{text "="}\<close>
lemma lfpI: "\<lbrakk>mono(f); a : f(lfp(f))\<rbrakk> \<Longrightarrow> a : lfp(f)"
apply (erule lfp_Tarski [THEN ssubst])
@@ -379,12 +379,12 @@
by simp
-ML {*
+ML \<open>
val coinduct3_tac = SUBPROOF (fn {context = ctxt, prems = mono :: prems, ...} =>
fast_tac (ctxt addIs (mono RS @{thm coinduct3_mono_lemma} RS @{thm lfpI}) :: prems) 1);
-*}
+\<close>
-method_setup coinduct3 = {* Scan.succeed (SIMPLE_METHOD' o coinduct3_tac) *}
+method_setup coinduct3 = \<open>Scan.succeed (SIMPLE_METHOD' o coinduct3_tac)\<close>
lemma ci3_RI: "\<lbrakk>mono(Agen); a : R\<rbrakk> \<Longrightarrow> a : lfp(\<lambda>x. Agen(x) Un R Un A)"
by coinduct3
@@ -396,21 +396,21 @@
lemma ci3_AI: "\<lbrakk>mono(Agen); a : A\<rbrakk> \<Longrightarrow> a : lfp(\<lambda>x. Agen(x) Un R Un A)"
by coinduct3
-ML {*
+ML \<open>
fun genIs_tac ctxt genXH gen_mono =
resolve_tac ctxt [genXH RS @{thm iffD2}] THEN'
simp_tac ctxt THEN'
TRY o fast_tac
(ctxt addIs [genXH RS @{thm iffD2}, gen_mono RS @{thm coinduct3_mono_lemma} RS @{thm lfpI}])
-*}
+\<close>
-method_setup genIs = {*
+method_setup genIs = \<open>
Attrib.thm -- Attrib.thm >>
(fn (genXH, gen_mono) => fn ctxt => SIMPLE_METHOD' (genIs_tac ctxt genXH gen_mono))
-*}
+\<close>
-subsection {* POgen *}
+subsection \<open>POgen\<close>
lemma PO_refl: "<a,a> : PO"
by (rule po_refl [THEN PO_iff [THEN iffD1]])
@@ -433,7 +433,7 @@
\<Longrightarrow> <h$t,h'$t'> : POgen(lfp(\<lambda>x. POgen(x) Un R Un PO))"
unfolding data_defs by (genIs POgenXH POgen_mono)+
-ML {*
+ML \<open>
fun POgen_tac ctxt (rla, rlb) i =
SELECT_GOAL (safe_tac ctxt) i THEN
resolve_tac ctxt [rlb RS (rla RS @{thm ssubst_pair})] i THEN
@@ -441,10 +441,10 @@
(@{thms POgenIs} @ [@{thm PO_refl} RS (@{thm POgen_mono} RS @{thm ci3_AI})] @
(@{thms POgenIs} RL [@{thm POgen_mono} RS @{thm ci3_AgenI}]) @
[@{thm POgen_mono} RS @{thm ci3_RI}]) i))
-*}
+\<close>
-subsection {* EQgen *}
+subsection \<open>EQgen\<close>
lemma EQ_refl: "<a,a> : EQ"
by (rule refl [THEN EQ_iff [THEN iffD1]])
@@ -467,7 +467,7 @@
\<Longrightarrow> <h$t,h'$t'> : EQgen(lfp(\<lambda>x. EQgen(x) Un R Un EQ))"
unfolding data_defs by (genIs EQgenXH EQgen_mono)+
-ML {*
+ML \<open>
fun EQgen_raw_tac ctxt i =
(REPEAT (resolve_tac ctxt (@{thms EQgenIs} @
[@{thm EQ_refl} RS (@{thm EQgen_mono} RS @{thm ci3_AI})] @
@@ -484,10 +484,10 @@
resolve_tac ctxt ((rews @ [@{thm refl}]) RL ((rews @ [@{thm refl}]) RL [@{thm ssubst_pair}])) i THEN
ALLGOALS (simp_tac ctxt) THEN
ALLGOALS (EQgen_raw_tac ctxt)) i
-*}
+\<close>
-method_setup EQgen = {*
+method_setup EQgen = \<open>
Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD' (EQgen_tac ctxt ths))
-*}
+\<close>
end
--- a/src/CCL/Wfd.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/CCL/Wfd.thy Thu Jul 23 14:25:05 2015 +0200
@@ -3,7 +3,7 @@
Copyright 1993 University of Cambridge
*)
-section {* Well-founded relations in CCL *}
+section \<open>Well-founded relations in CCL\<close>
theory Wfd
imports Trancl Type Hered
@@ -46,12 +46,12 @@
apply blast
done
-method_setup wfd_strengthen = {*
+method_setup wfd_strengthen = \<open>
Scan.lift Args.name_inner_syntax >> (fn s => fn ctxt =>
SIMPLE_METHOD' (fn i =>
Rule_Insts.res_inst_tac ctxt [((("Q", 0), Position.none), s)] [] @{thm wfd_strengthen_lemma} i
THEN assume_tac ctxt (i + 1)))
-*}
+\<close>
lemma wf_anti_sym: "\<lbrakk>Wfd(r); <a,x>:r; <x,a>:r\<rbrakk> \<Longrightarrow> P"
apply (subgoal_tac "ALL x. <a,x>:r \<longrightarrow> <x,a>:r \<longrightarrow> P")
@@ -66,7 +66,7 @@
done
-subsection {* Irreflexive transitive closure *}
+subsection \<open>Irreflexive transitive closure\<close>
lemma trancl_wf:
assumes 1: "Wfd(R)"
@@ -85,7 +85,7 @@
done
-subsection {* Lexicographic Ordering *}
+subsection \<open>Lexicographic Ordering\<close>
lemma lexXH:
"p : ra**rb \<longleftrightarrow> (EX a a' b b'. p = <<a,b>,<a',b'>> \<and> (<a,a'> : ra | a=a' \<and> <b,b'> : rb))"
@@ -128,7 +128,7 @@
done
-subsection {* Mapping *}
+subsection \<open>Mapping\<close>
lemma wmapXH: "p : wmap(f,r) \<longleftrightarrow> (EX x y. p=<x,y> \<and> <f(x),f(y)> : r)"
unfolding wmap_def by blast
@@ -156,7 +156,7 @@
done
-subsection {* Projections *}
+subsection \<open>Projections\<close>
lemma wfstI: "<xa,ya> : r \<Longrightarrow> <<xa,xb>,<ya,yb>> : wmap(fst,r)"
apply (rule wmapI)
@@ -174,7 +174,7 @@
done
-subsection {* Ground well-founded relations *}
+subsection \<open>Ground well-founded relations\<close>
lemma wfI: "\<lbrakk>Wfd(r); a : r\<rbrakk> \<Longrightarrow> a : wf(r)"
unfolding wf_def by blast
@@ -220,7 +220,7 @@
done
-subsection {* General Recursive Functions *}
+subsection \<open>General Recursive Functions\<close>
lemma letrecT:
assumes 1: "a : A"
@@ -282,7 +282,7 @@
lemmas letrecTs = letrecT letrec2T letrec3T
-subsection {* Type Checking for Recursive Calls *}
+subsection \<open>Type Checking for Recursive Calls\<close>
lemma rcallT:
"\<lbrakk>ALL x:{x:A.<x,p>:wf(R)}.g(x):D(x);
@@ -303,7 +303,7 @@
lemmas rcallTs = rcallT rcall2T rcall3T
-subsection {* Instantiating an induction hypothesis with an equality assumption *}
+subsection \<open>Instantiating an induction hypothesis with an equality assumption\<close>
lemma hyprcallT:
assumes 1: "g(a) = b"
@@ -360,7 +360,7 @@
lemmas hyprcallTs = hyprcallT hyprcall2T hyprcall3T
-subsection {* Rules to Remove Induction Hypotheses after Type Checking *}
+subsection \<open>Rules to Remove Induction Hypotheses after Type Checking\<close>
lemma rmIH1: "\<lbrakk>ALL x:{x:A.<x,p>:wf(R)}.g(x):D(x); P\<rbrakk> \<Longrightarrow> P" .
@@ -372,7 +372,7 @@
lemmas rmIHs = rmIH1 rmIH2 rmIH3
-subsection {* Lemmas for constructors and subtypes *}
+subsection \<open>Lemmas for constructors and subtypes\<close>
(* 0-ary constructors do not need additional rules as they are handled *)
(* correctly by applying SubtypeI *)
@@ -404,9 +404,9 @@
lemmas rcall_lemmas = asm_rl rcall_lemma1 SubtypeD1 rcall_lemma2
-subsection {* Typechecking *}
+subsection \<open>Typechecking\<close>
-ML {*
+ML \<open>
local
val type_rls =
@@ -478,35 +478,35 @@
SELECT_GOAL (REPEAT_FIRST (tc_step_tac ctxt rls) THEN clean_ccs_tac ctxt) i
end
-*}
+\<close>
-method_setup typechk = {*
+method_setup typechk = \<open>
Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD' (typechk_tac ctxt ths))
-*}
+\<close>
-method_setup clean_ccs = {*
+method_setup clean_ccs = \<open>
Scan.succeed (SIMPLE_METHOD o clean_ccs_tac)
-*}
+\<close>
-method_setup gen_ccs = {*
+method_setup gen_ccs = \<open>
Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD' (gen_ccs_tac ctxt ths))
-*}
+\<close>
-subsection {* Evaluation *}
+subsection \<open>Evaluation\<close>
named_theorems eval "evaluation rules"
-ML {*
+ML \<open>
fun eval_tac ths =
Subgoal.FOCUS_PREMS (fn {context = ctxt, prems, ...} =>
let val eval_rules = Named_Theorems.get ctxt @{named_theorems eval}
in DEPTH_SOLVE_1 (resolve_tac ctxt (ths @ prems @ rev eval_rules) 1) end)
-*}
+\<close>
-method_setup eval = {*
+method_setup eval = \<open>
Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD' (CHANGED o eval_tac ths ctxt))
-*}
+\<close>
lemmas eval_rls [eval] = trueV falseV pairV lamV caseVtrue caseVfalse caseVpair caseVlam
@@ -523,9 +523,9 @@
shows "let x be t in f(x) ---> c"
apply (unfold let_def)
apply (rule 1 [THEN canonical])
- apply (tactic {*
+ apply (tactic \<open>
REPEAT (DEPTH_SOLVE_1 (resolve_tac @{context} (@{thms assms} @ @{thms eval_rls}) 1 ORELSE
- eresolve_tac @{context} @{thms substitute} 1)) *})
+ eresolve_tac @{context} @{thms substitute} 1))\<close>)
done
lemma fixV: "f(fix(f)) ---> c \<Longrightarrow> fix(f) ---> c"
@@ -566,7 +566,7 @@
unfolding data_defs by eval+
-subsection {* Factorial *}
+subsection \<open>Factorial\<close>
schematic_lemma
"letrec f n be ncase(n,succ(zero),\<lambda>x. nrec(n,zero,\<lambda>y g. nrec(f(x),g,\<lambda>z h. succ(h))))
@@ -578,7 +578,7 @@
in f(succ(succ(succ(zero)))) ---> ?a"
by eval
-subsection {* Less Than Or Equal *}
+subsection \<open>Less Than Or Equal\<close>
schematic_lemma
"letrec f p be split(p,\<lambda>m n. ncase(m,true,\<lambda>x. ncase(n,false,\<lambda>y. f(<x,y>))))
@@ -596,7 +596,7 @@
by eval
-subsection {* Reverse *}
+subsection \<open>Reverse\<close>
schematic_lemma
"letrec id l be lcase(l,[],\<lambda>x xs. x$id(xs))
--- a/src/CCL/ex/Flag.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/CCL/ex/Flag.thy Thu Jul 23 14:25:05 2015 +0200
@@ -3,8 +3,8 @@
Copyright 1993 University of Cambridge
*)
-section {* Dutch national flag program -- except that the point of Dijkstra's example was to use
- arrays and this uses lists. *}
+section \<open>Dutch national flag program -- except that the point of Dijkstra's example was to use
+ arrays and this uses lists.\<close>
theory Flag
imports List
--- a/src/CCL/ex/List.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/CCL/ex/List.thy Thu Jul 23 14:25:05 2015 +0200
@@ -3,7 +3,7 @@
Copyright 1993 University of Cambridge
*)
-section {* Programs defined over lists *}
+section \<open>Programs defined over lists\<close>
theory List
imports Nat
--- a/src/CCL/ex/Nat.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/CCL/ex/Nat.thy Thu Jul 23 14:25:05 2015 +0200
@@ -3,7 +3,7 @@
Copyright 1993 University of Cambridge
*)
-section {* Programs defined over the natural numbers *}
+section \<open>Programs defined over the natural numbers\<close>
theory Nat
imports "../Wfd"
@@ -96,7 +96,7 @@
done
-subsection {* Termination Conditions for Ackermann's Function *}
+subsection \<open>Termination Conditions for Ackermann's Function\<close>
lemmas relI = NatPR_wf [THEN NatPR_wf [THEN lex_wf, THEN wfI]]
--- a/src/CCL/ex/Stream.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/CCL/ex/Stream.thy Thu Jul 23 14:25:05 2015 +0200
@@ -3,7 +3,7 @@
Copyright 1993 University of Cambridge
*)
-section {* Programs defined over streams *}
+section \<open>Programs defined over streams\<close>
theory Stream
imports List
@@ -22,7 +22,7 @@
*)
-subsection {* Map of composition is composition of maps *}
+subsection \<open>Map of composition is composition of maps\<close>
lemma map_comp:
assumes 1: "l:Lists(A)"
@@ -49,7 +49,7 @@
done
-subsection {* Mapping distributes over append *}
+subsection \<open>Mapping distributes over append\<close>
lemma map_append:
assumes "l:Lists(A)"
@@ -67,7 +67,7 @@
done
-subsection {* Append is associative *}
+subsection \<open>Append is associative\<close>
lemma append_assoc:
assumes "k:Lists(A)"
@@ -82,12 +82,12 @@
apply EQgen
prefer 2
apply blast
- apply (tactic {* DEPTH_SOLVE (eresolve_tac @{context} [XH_to_E @{thm ListsXH}] 1
- THEN EQgen_tac @{context} [] 1) *})
+ apply (tactic \<open>DEPTH_SOLVE (eresolve_tac @{context} [XH_to_E @{thm ListsXH}] 1
+ THEN EQgen_tac @{context} [] 1)\<close>)
done
-subsection {* Appending anything to an infinite list doesn't alter it *}
+subsection \<open>Appending anything to an infinite list doesn't alter it\<close>
lemma ilist_append:
assumes "l:ILists(A)"
--- a/src/CTT/Arith.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/CTT/Arith.thy Thu Jul 23 14:25:05 2015 +0200
@@ -3,13 +3,13 @@
Copyright 1991 University of Cambridge
*)
-section {* Elementary arithmetic *}
+section \<open>Elementary arithmetic\<close>
theory Arith
imports Bool
begin
-subsection {* Arithmetic operators and their definitions *}
+subsection \<open>Arithmetic operators and their definitions\<close>
definition
add :: "[i,i]\<Rightarrow>i" (infixr "#+" 65) where
@@ -46,7 +46,7 @@
lemmas arith_defs = add_def diff_def absdiff_def mult_def mod_def div_def
-subsection {* Proofs about elementary arithmetic: addition, multiplication, etc. *}
+subsection \<open>Proofs about elementary arithmetic: addition, multiplication, etc.\<close>
(** Addition *)
@@ -144,7 +144,7 @@
done
-subsection {* Simplification *}
+subsection \<open>Simplification\<close>
lemmas arith_typing_rls = add_typing mult_typing diff_typing
and arith_congr_rls = add_typingL mult_typingL diff_typingL
@@ -155,7 +155,7 @@
multC0 multC_succ
diffC0 diff_0_eq_0 diff_succ_succ
-ML {*
+ML \<open>
structure Arith_simp_data: TSIMP_DATA =
struct
@@ -180,18 +180,18 @@
(Arith_simp.cond_norm_tac ctxt (prove_cond_tac ctxt, congr_rls, prems))
end
-*}
+\<close>
-method_setup arith_rew = {*
+method_setup arith_rew = \<open>
Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD (arith_rew_tac ctxt ths))
-*}
+\<close>
-method_setup hyp_arith_rew = {*
+method_setup hyp_arith_rew = \<open>
Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD (hyp_arith_rew_tac ctxt ths))
-*}
+\<close>
-subsection {* Addition *}
+subsection \<open>Addition\<close>
(*Associative law for addition*)
lemma add_assoc: "\<lbrakk>a:N; b:N; c:N\<rbrakk> \<Longrightarrow> (a #+ b) #+ c = a #+ (b #+ c) : N"
@@ -214,7 +214,7 @@
done
-subsection {* Multiplication *}
+subsection \<open>Multiplication\<close>
(*right annihilation in product*)
lemma mult_0_right: "a:N \<Longrightarrow> a #* 0 = 0 : N"
@@ -248,11 +248,11 @@
done
-subsection {* Difference *}
+subsection \<open>Difference\<close>
-text {*
+text \<open>
Difference on natural numbers, without negative numbers
- a - b = 0 iff a<=b a - b = succ(c) iff a>b *}
+ a - b = 0 iff a<=b a - b = succ(c) iff a>b\<close>
lemma diff_self_eq_0: "a:N \<Longrightarrow> a - a = 0 : N"
apply (NE a)
@@ -299,7 +299,7 @@
done
-subsection {* Absolute difference *}
+subsection \<open>Absolute difference\<close>
(*typing of absolute difference: short and long versions*)
@@ -380,7 +380,7 @@
done
-subsection {* Remainder and Quotient *}
+subsection \<open>Remainder and Quotient\<close>
(*typing of remainder: short and long versions*)
--- a/src/CTT/Bool.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/CTT/Bool.thy Thu Jul 23 14:25:05 2015 +0200
@@ -3,7 +3,7 @@
Copyright 1991 University of Cambridge
*)
-section {* The two-element type (booleans and conditionals) *}
+section \<open>The two-element type (booleans and conditionals)\<close>
theory Bool
imports CTT
@@ -28,7 +28,7 @@
lemmas bool_defs = Bool_def true_def false_def cond_def
-subsection {* Derivation of rules for the type Bool *}
+subsection \<open>Derivation of rules for the type Bool\<close>
(*formation rule*)
lemma boolF: "Bool type"
--- a/src/CTT/CTT.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/CTT/CTT.thy Thu Jul 23 14:25:05 2015 +0200
@@ -3,7 +3,7 @@
Copyright 1993 University of Cambridge
*)
-section {* Constructive Type Theory *}
+section \<open>Constructive Type Theory\<close>
theory CTT
imports Pure
@@ -314,9 +314,9 @@
done
-subsection {* Tactics for type checking *}
+subsection \<open>Tactics for type checking\<close>
-ML {*
+ML \<open>
local
@@ -336,13 +336,13 @@
end;
-*}
+\<close>
(*For simplification: type formation and checking,
but no equalities between terms*)
lemmas routine_rls = form_rls formL_rls refl_type element_rls
-ML {*
+ML \<open>
local
val equal_rls = @{thms form_rls} @ @{thms element_rls} @ @{thms intrL_rls} @
@{thms elimL_rls} @ @{thms refl_elem}
@@ -378,15 +378,15 @@
(ASSUME ctxt (filt_resolve_from_net_tac ctxt 3 (Tactic.build_net (thms @ equal_rls))))
end
-*}
+\<close>
-method_setup form = {* Scan.succeed (fn ctxt => SIMPLE_METHOD (form_tac ctxt)) *}
-method_setup typechk = {* Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD (typechk_tac ctxt ths)) *}
-method_setup intr = {* Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD (intr_tac ctxt ths)) *}
-method_setup equal = {* Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD (equal_tac ctxt ths)) *}
+method_setup form = \<open>Scan.succeed (fn ctxt => SIMPLE_METHOD (form_tac ctxt))\<close>
+method_setup typechk = \<open>Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD (typechk_tac ctxt ths))\<close>
+method_setup intr = \<open>Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD (intr_tac ctxt ths))\<close>
+method_setup equal = \<open>Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD (equal_tac ctxt ths))\<close>
-subsection {* Simplification *}
+subsection \<open>Simplification\<close>
(*To simplify the type in a goal*)
lemma replace_type: "\<lbrakk>B = A; a : A\<rbrakk> \<Longrightarrow> a : B"
@@ -409,7 +409,7 @@
(*Simplification rules for Constructive Type Theory*)
lemmas reduction_rls = comp_rls [THEN trans_elem]
-ML {*
+ML \<open>
(*Converts each goal "e : Eq(A,a,b)" into "a=b:A" for simplification.
Uses other intro rules to avoid changing flexible goals.*)
val eqintr_net = Tactic.build_net @{thms EqI intr_rls}
@@ -466,21 +466,21 @@
(*Fails unless it solves the goal!*)
fun pc_tac ctxt thms = DEPTH_SOLVE_1 o (step_tac ctxt thms)
-*}
+\<close>
-method_setup eqintr = {* Scan.succeed (SIMPLE_METHOD o eqintr_tac) *}
-method_setup NE = {*
+method_setup eqintr = \<open>Scan.succeed (SIMPLE_METHOD o eqintr_tac)\<close>
+method_setup NE = \<open>
Scan.lift Args.name_inner_syntax >> (fn s => fn ctxt => SIMPLE_METHOD' (NE_tac ctxt s))
-*}
-method_setup pc = {* Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD' (pc_tac ctxt ths)) *}
-method_setup add_mp = {* Scan.succeed (SIMPLE_METHOD' o add_mp_tac) *}
+\<close>
+method_setup pc = \<open>Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD' (pc_tac ctxt ths))\<close>
+method_setup add_mp = \<open>Scan.succeed (SIMPLE_METHOD' o add_mp_tac)\<close>
ML_file "rew.ML"
-method_setup rew = {* Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD (rew_tac ctxt ths)) *}
-method_setup hyp_rew = {* Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD (hyp_rew_tac ctxt ths)) *}
+method_setup rew = \<open>Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD (rew_tac ctxt ths))\<close>
+method_setup hyp_rew = \<open>Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD (hyp_rew_tac ctxt ths))\<close>
-subsection {* The elimination rules for fst/snd *}
+subsection \<open>The elimination rules for fst/snd\<close>
lemma SumE_fst: "p : Sum(A,B) \<Longrightarrow> fst(p) : A"
apply (unfold basic_defs)
--- a/src/CTT/Main.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/CTT/Main.thy Thu Jul 23 14:25:05 2015 +0200
@@ -1,4 +1,4 @@
-section {* Main includes everything *}
+section \<open>Main includes everything\<close>
theory Main
imports CTT Arith Bool
--- a/src/CTT/ex/Elimination.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/CTT/ex/Elimination.thy Thu Jul 23 14:25:05 2015 +0200
@@ -182,7 +182,7 @@
and "\<And>z. z:A*B \<Longrightarrow> C(z) type"
shows "?a : (SUM z:A*B. C(z)) --> (SUM u:A. SUM v:B. C(<u,v>))"
apply (rule intr_rls)
-apply (tactic {* biresolve_tac @{context} safe_brls 2 *})
+apply (tactic \<open>biresolve_tac @{context} safe_brls 2\<close>)
(*Now must convert assumption C(z) into antecedent C(<kd,ke>) *)
apply (rule_tac [2] a = "y" in ProdE)
apply (typechk assms)
--- a/src/CTT/ex/Equality.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/CTT/ex/Equality.thy Thu Jul 23 14:25:05 2015 +0200
@@ -47,7 +47,7 @@
lemma "p : Sum(A,B) \<Longrightarrow> <split(p,\<lambda>x y. x), split(p,\<lambda>x y. y)> = p : Sum(A,B)"
apply (rule EqE)
apply (rule elim_rls, assumption)
-apply (tactic {* DEPTH_SOLVE_1 (rew_tac @{context} []) *}) (*!!!!!!!*)
+apply (tactic \<open>DEPTH_SOLVE_1 (rew_tac @{context} [])\<close>) (*!!!!!!!*)
done
lemma "\<lbrakk>a : A; b : B\<rbrakk> \<Longrightarrow> (lam u. split(u, \<lambda>v w.<w,v>)) ` <a,b> = <b,a> : SUM x:B. A"
--- a/src/CTT/ex/Typechecking.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/CTT/ex/Typechecking.thy Thu Jul 23 14:25:05 2015 +0200
@@ -9,7 +9,7 @@
imports "../CTT"
begin
-subsection {* Single-step proofs: verifying that a type is well-formed *}
+subsection \<open>Single-step proofs: verifying that a type is well-formed\<close>
schematic_lemma "?A type"
apply (rule form_rls)
@@ -31,7 +31,7 @@
done
-subsection {* Multi-step proofs: Type inference *}
+subsection \<open>Multi-step proofs: Type inference\<close>
lemma "PROD w:N. N + N type"
apply form
@@ -67,7 +67,7 @@
(*Proofs involving arbitrary types.
For concreteness, every type variable left over is forced to be N*)
method_setup N =
- {* Scan.succeed (fn ctxt => SIMPLE_METHOD (TRYALL (resolve_tac ctxt @{thms NF}))) *}
+ \<open>Scan.succeed (fn ctxt => SIMPLE_METHOD (TRYALL (resolve_tac ctxt @{thms NF})))\<close>
schematic_lemma "lam w. <w,w> : ?A"
apply typechk
--- a/src/FOL/FOL.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/FOL/FOL.thy Thu Jul 23 14:25:05 2015 +0200
@@ -2,7 +2,7 @@
Author: Lawrence C Paulson and Markus Wenzel
*)
-section {* Classical first-order logic *}
+section \<open>Classical first-order logic\<close>
theory FOL
imports IFOL
@@ -14,13 +14,13 @@
ML_file "~~/src/Provers/clasimp.ML"
-subsection {* The classical axiom *}
+subsection \<open>The classical axiom\<close>
axiomatization where
classical: "(~P ==> P) ==> P"
-subsection {* Lemmas and proof tools *}
+subsection \<open>Lemmas and proof tools\<close>
lemma ccontr: "(\<not> P \<Longrightarrow> False) \<Longrightarrow> P"
by (erule FalseE [THEN classical])
@@ -65,15 +65,15 @@
apply (erule r1)
done
-ML {*
+ML \<open>
fun case_tac ctxt a fixes =
Rule_Insts.res_inst_tac ctxt [((("P", 0), Position.none), a)] fixes @{thm case_split}
-*}
+\<close>
-method_setup case_tac = {*
+method_setup case_tac = \<open>
Args.goal_spec -- Scan.lift (Args.name_inner_syntax -- Parse.for_fixes) >>
(fn (quant, (s, fixes)) => fn ctxt => SIMPLE_METHOD'' quant (case_tac ctxt s fixes))
-*} "case_tac emulation (dynamic instantiation!)"
+\<close> "case_tac emulation (dynamic instantiation!)"
(*** Special elimination rules *)
@@ -163,9 +163,9 @@
by (rule classical) iprover
-section {* Classical Reasoner *}
+section \<open>Classical Reasoner\<close>
-ML {*
+ML \<open>
structure Cla = Classical
(
val imp_elim = @{thm imp_elim}
@@ -178,21 +178,21 @@
structure Basic_Classical: BASIC_CLASSICAL = Cla;
open Basic_Classical;
-*}
+\<close>
(*Propositional rules*)
lemmas [intro!] = refl TrueI conjI disjCI impI notI iffI
and [elim!] = conjE disjE impCE FalseE iffCE
-ML {* val prop_cs = claset_of @{context} *}
+ML \<open>val prop_cs = claset_of @{context}\<close>
(*Quantifier rules*)
lemmas [intro!] = allI ex_ex1I
and [intro] = exI
and [elim!] = exE alt_ex1E
and [elim] = allE
-ML {* val FOL_cs = claset_of @{context} *}
+ML \<open>val FOL_cs = claset_of @{context}\<close>
-ML {*
+ML \<open>
structure Blast = Blast
(
structure Classical = Cla
@@ -204,7 +204,7 @@
val hyp_subst_tac = Hypsubst.blast_hyp_subst_tac
);
val blast_tac = Blast.blast_tac;
-*}
+\<close>
lemma ex1_functional: "[| EX! z. P(a,z); P(a,b); P(a,c) |] ==> b = c"
@@ -320,10 +320,10 @@
ML_file "simpdata.ML"
-simproc_setup defined_Ex ("EX x. P(x)") = {* fn _ => Quantifier1.rearrange_ex *}
-simproc_setup defined_All ("ALL x. P(x)") = {* fn _ => Quantifier1.rearrange_all *}
+simproc_setup defined_Ex ("EX x. P(x)") = \<open>fn _ => Quantifier1.rearrange_ex\<close>
+simproc_setup defined_All ("ALL x. P(x)") = \<open>fn _ => Quantifier1.rearrange_all\<close>
-ML {*
+ML \<open>
(*intuitionistic simprules only*)
val IFOL_ss =
put_simpset FOL_basic_ss @{context}
@@ -337,17 +337,17 @@
put_simpset IFOL_ss @{context}
addsimps @{thms cla_simps cla_ex_simps cla_all_simps}
|> simpset_of;
-*}
+\<close>
-setup {*
+setup \<open>
map_theory_simpset (put_simpset FOL_ss) #>
Simplifier.method_setup Splitter.split_modifiers
-*}
+\<close>
ML_file "~~/src/Tools/eqsubst.ML"
-subsection {* Other simple lemmas *}
+subsection \<open>Other simple lemmas\<close>
lemma [simp]: "((P-->R) <-> (Q-->R)) <-> ((P<->Q) | R)"
by blast
@@ -380,9 +380,9 @@
by blast
-subsection {* Proof by cases and induction *}
+subsection \<open>Proof by cases and induction\<close>
-text {* Proper handling of non-atomic rule statements. *}
+text \<open>Proper handling of non-atomic rule statements.\<close>
context
begin
@@ -409,10 +409,10 @@
lemmas induct_rulify_fallback =
induct_forall_def induct_implies_def induct_equal_def induct_conj_def
-text {* Method setup. *}
+text \<open>Method setup.\<close>
ML_file "~~/src/Tools/induct.ML"
-ML {*
+ML \<open>
structure Induct = Induct
(
val cases_default = @{thm case_split}
@@ -423,7 +423,7 @@
fun dest_def _ = NONE
fun trivial_tac _ _ = no_tac
);
-*}
+\<close>
declare case_split [cases type: o]
--- a/src/FOL/IFOL.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/FOL/IFOL.thy Thu Jul 23 14:25:05 2015 +0200
@@ -2,7 +2,7 @@
Author: Lawrence C Paulson and Markus Wenzel
*)
-section {* Intuitionistic first-order logic *}
+section \<open>Intuitionistic first-order logic\<close>
theory IFOL
imports Pure
@@ -20,7 +20,7 @@
ML_file "~~/src/Tools/atomize_elim.ML"
-subsection {* Syntax and axiomatic basis *}
+subsection \<open>Syntax and axiomatic basis\<close>
setup Pure_Thy.old_appl_syntax_setup
@@ -33,7 +33,7 @@
Trueprop :: "o => prop" ("(_)" 5)
-subsubsection {* Equality *}
+subsubsection \<open>Equality\<close>
axiomatization
eq :: "['a, 'a] => o" (infixl "=" 50)
@@ -42,7 +42,7 @@
subst: "a=b \<Longrightarrow> P(a) \<Longrightarrow> P(b)"
-subsubsection {* Propositional logic *}
+subsubsection \<open>Propositional logic\<close>
axiomatization
False :: o and
@@ -64,7 +64,7 @@
FalseE: "False ==> P"
-subsubsection {* Quantifiers *}
+subsubsection \<open>Quantifiers\<close>
axiomatization
All :: "('a => o) => o" (binder "ALL " 10) and
@@ -76,7 +76,7 @@
exE: "[| EX x. P(x); !!x. P(x) ==> R |] ==> R"
-subsubsection {* Definitions *}
+subsubsection \<open>Definitions\<close>
definition "True == False-->False"
definition Not ("~ _" [40] 40) where not_def: "~P == P-->False"
@@ -85,12 +85,12 @@
definition Ex1 :: "('a => o) => o" (binder "EX! " 10)
where ex1_def: "EX! x. P(x) == EX x. P(x) & (ALL y. P(y) --> y=x)"
-axiomatization where -- {* Reflection, admissible *}
+axiomatization where -- \<open>Reflection, admissible\<close>
eq_reflection: "(x=y) ==> (x==y)" and
iff_reflection: "(P<->Q) ==> (P==Q)"
-subsubsection {* Additional notation *}
+subsubsection \<open>Additional notation\<close>
abbreviation not_equal :: "['a, 'a] => o" (infixl "~=" 50)
where "x ~= y == ~ (x = y)"
@@ -120,7 +120,7 @@
Ex1 (binder "\<exists>!" 10)
-subsection {* Lemmas and proof tools *}
+subsection \<open>Lemmas and proof tools\<close>
lemmas strip = impI allI
@@ -146,7 +146,7 @@
shows R
apply (rule r)
apply (rule major [THEN mp])
- apply (rule `P`)
+ apply (rule \<open>P\<close>)
done
lemma allE:
@@ -186,7 +186,7 @@
shows Q
apply (rule r)
apply (rule impI)
- apply (erule notE [OF `~P`])
+ apply (erule notE [OF \<open>~P\<close>])
done
(* For substitution into an assumption P, reduce Q to P-->Q, substitute into
@@ -207,10 +207,10 @@
(*** Modus Ponens Tactics ***)
(*Finds P-->Q and P in the assumptions, replaces implication by Q *)
-ML {*
+ML \<open>
fun mp_tac ctxt i = eresolve_tac ctxt @{thms notE impE} i THEN assume_tac ctxt i
fun eq_mp_tac ctxt i = eresolve_tac ctxt @{thms notE impE} i THEN eq_assume_tac i
-*}
+\<close>
(*** If-and-only-if ***)
@@ -303,11 +303,11 @@
(*** <-> congruence rules for simplification ***)
(*Use iffE on a premise. For conj_cong, imp_cong, all_cong, ex_cong*)
-ML {*
+ML \<open>
fun iff_tac ctxt prems i =
resolve_tac ctxt (prems RL @{thms iffE}) i THEN
REPEAT1 (eresolve_tac ctxt @{thms asm_rl mp} i)
-*}
+\<close>
method_setup iff =
\<open>Attrib.thms >> (fn prems => fn ctxt => SIMPLE_METHOD' (iff_tac ctxt prems))\<close>
@@ -558,20 +558,20 @@
apply (rule disjI2) apply assumption
done
-ML {*
+ML \<open>
structure Project_Rule = Project_Rule
(
val conjunct1 = @{thm conjunct1}
val conjunct2 = @{thm conjunct2}
val mp = @{thm mp}
)
-*}
+\<close>
ML_file "fologic.ML"
lemma thin_refl: "[|x=x; PROP W|] ==> PROP W" .
-ML {*
+ML \<open>
structure Hypsubst = Hypsubst
(
val dest_eq = FOLogic.dest_eq
@@ -586,14 +586,14 @@
val thin_refl = @{thm thin_refl}
);
open Hypsubst;
-*}
+\<close>
ML_file "intprover.ML"
-subsection {* Intuitionistic Reasoning *}
+subsection \<open>Intuitionistic Reasoning\<close>
-setup {* Intuitionistic.method_setup @{binding iprover} *}
+setup \<open>Intuitionistic.method_setup @{binding iprover}\<close>
lemma impE':
assumes 1: "P --> Q"
@@ -629,7 +629,7 @@
and [Pure.elim 2] = allE notE' impE'
and [Pure.intro] = exI disjI2 disjI1
-setup {* Context_Rules.addSWrapper (fn ctxt => fn tac => hyp_subst_tac ctxt ORELSE' tac) *}
+setup \<open>Context_Rules.addSWrapper (fn ctxt => fn tac => hyp_subst_tac ctxt ORELSE' tac)\<close>
lemma iff_not_sym: "~ (Q <-> P) ==> ~ (P <-> Q)"
@@ -645,7 +645,7 @@
done
-subsection {* Atomizing meta-level rules *}
+subsection \<open>Atomizing meta-level rules\<close>
lemma atomize_all [atomize]: "(!!x. P(x)) == Trueprop (ALL x. P(x))"
proof
@@ -668,7 +668,7 @@
lemma atomize_eq [atomize]: "(x == y) == Trueprop (x = y)"
proof
assume "x == y"
- show "x = y" unfolding `x == y` by (rule refl)
+ show "x = y" unfolding \<open>x == y\<close> by (rule refl)
next
assume "x = y"
then show "x == y" by (rule eq_reflection)
@@ -677,7 +677,7 @@
lemma atomize_iff [atomize]: "(A == B) == Trueprop (A <-> B)"
proof
assume "A == B"
- show "A <-> B" unfolding `A == B` by (rule iff_refl)
+ show "A <-> B" unfolding \<open>A == B\<close> by (rule iff_refl)
next
assume "A <-> B"
then show "A == B" by (rule iff_reflection)
@@ -704,7 +704,7 @@
and [symmetric, defn] = atomize_all atomize_imp atomize_eq atomize_iff
-subsection {* Atomizing elimination rules *}
+subsection \<open>Atomizing elimination rules\<close>
lemma atomize_exL[atomize_elim]: "(!!x. P(x) ==> Q) == ((EX x. P(x)) ==> Q)"
by rule iprover+
@@ -718,7 +718,7 @@
lemma atomize_elimL[atomize_elim]: "(!!B. (A ==> B) ==> B) == Trueprop(A)" ..
-subsection {* Calculational rules *}
+subsection \<open>Calculational rules\<close>
lemma forw_subst: "a = b ==> P(b) ==> P(a)"
by (rule ssubst)
@@ -726,9 +726,9 @@
lemma back_subst: "P(a) ==> a = b ==> P(b)"
by (rule subst)
-text {*
+text \<open>
Note that this list of rules is in reverse order of priorities.
-*}
+\<close>
lemmas basic_trans_rules [trans] =
forw_subst
@@ -737,7 +737,7 @@
mp
trans
-subsection {* ``Let'' declarations *}
+subsection \<open>``Let'' declarations\<close>
nonterminal letbinds and letbind
@@ -763,7 +763,7 @@
done
-subsection {* Intuitionistic simplification rules *}
+subsection \<open>Intuitionistic simplification rules\<close>
lemma conj_simps:
"P & True <-> P"
@@ -830,7 +830,7 @@
by iprover+
-text {* Conversion into rewrite rules *}
+text \<open>Conversion into rewrite rules\<close>
lemma P_iff_F: "~P ==> (P <-> False)" by iprover
lemma iff_reflection_F: "~P ==> (P == False)" by (rule P_iff_F [THEN iff_reflection])
@@ -839,7 +839,7 @@
lemma iff_reflection_T: "P ==> (P == True)" by (rule P_iff_T [THEN iff_reflection])
-text {* More rewrite rules *}
+text \<open>More rewrite rules\<close>
lemma conj_commute: "P&Q <-> Q&P" by iprover
lemma conj_left_commute: "P&(Q&R) <-> Q&(P&R)" by iprover
--- a/src/FOL/ex/Classical.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/FOL/ex/Classical.thy Thu Jul 23 14:25:05 2015 +0200
@@ -3,14 +3,14 @@
Copyright 1994 University of Cambridge
*)
-section{*Classical Predicate Calculus Problems*}
+section\<open>Classical Predicate Calculus Problems\<close>
theory Classical imports FOL begin
lemma "(P --> Q | R) --> (P-->Q) | (P-->R)"
by blast
-text{*If and only if*}
+text\<open>If and only if\<close>
lemma "(P<->Q) <-> (Q<->P)"
by blast
@@ -19,7 +19,7 @@
by blast
-text{*Sample problems from
+text\<open>Sample problems from
F. J. Pelletier,
Seventy-Five Problems for Testing Automatic Theorem Provers,
J. Automated Reasoning 2 (1986), 191-216.
@@ -27,79 +27,79 @@
The hardest problems -- judging by experience with several theorem provers,
including matrix ones -- are 34 and 43.
-*}
+\<close>
-subsection{*Pelletier's examples*}
+subsection\<open>Pelletier's examples\<close>
-text{*1*}
+text\<open>1\<close>
lemma "(P-->Q) <-> (~Q --> ~P)"
by blast
-text{*2*}
+text\<open>2\<close>
lemma "~ ~ P <-> P"
by blast
-text{*3*}
+text\<open>3\<close>
lemma "~(P-->Q) --> (Q-->P)"
by blast
-text{*4*}
+text\<open>4\<close>
lemma "(~P-->Q) <-> (~Q --> P)"
by blast
-text{*5*}
+text\<open>5\<close>
lemma "((P|Q)-->(P|R)) --> (P|(Q-->R))"
by blast
-text{*6*}
+text\<open>6\<close>
lemma "P | ~ P"
by blast
-text{*7*}
+text\<open>7\<close>
lemma "P | ~ ~ ~ P"
by blast
-text{*8. Peirce's law*}
+text\<open>8. Peirce's law\<close>
lemma "((P-->Q) --> P) --> P"
by blast
-text{*9*}
+text\<open>9\<close>
lemma "((P|Q) & (~P|Q) & (P| ~Q)) --> ~ (~P | ~Q)"
by blast
-text{*10*}
+text\<open>10\<close>
lemma "(Q-->R) & (R-->P&Q) & (P-->Q|R) --> (P<->Q)"
by blast
-text{*11. Proved in each direction (incorrectly, says Pelletier!!) *}
+text\<open>11. Proved in each direction (incorrectly, says Pelletier!!)\<close>
lemma "P<->P"
by blast
-text{*12. "Dijkstra's law"*}
+text\<open>12. "Dijkstra's law"\<close>
lemma "((P <-> Q) <-> R) <-> (P <-> (Q <-> R))"
by blast
-text{*13. Distributive law*}
+text\<open>13. Distributive law\<close>
lemma "P | (Q & R) <-> (P | Q) & (P | R)"
by blast
-text{*14*}
+text\<open>14\<close>
lemma "(P <-> Q) <-> ((Q | ~P) & (~Q|P))"
by blast
-text{*15*}
+text\<open>15\<close>
lemma "(P --> Q) <-> (~P | Q)"
by blast
-text{*16*}
+text\<open>16\<close>
lemma "(P-->Q) | (Q-->P)"
by blast
-text{*17*}
+text\<open>17\<close>
lemma "((P & (Q-->R))-->S) <-> ((~P | Q | S) & (~P | ~R | S))"
by blast
-subsection{*Classical Logic: examples with quantifiers*}
+subsection\<open>Classical Logic: examples with quantifiers\<close>
lemma "(\<forall>x. P(x) & Q(x)) <-> (\<forall>x. P(x)) & (\<forall>x. Q(x))"
by blast
@@ -113,23 +113,23 @@
lemma "(\<forall>x. P(x)) | Q <-> (\<forall>x. P(x) | Q)"
by blast
-text{*Discussed in Avron, Gentzen-Type Systems, Resolution and Tableaux,
- JAR 10 (265-281), 1993. Proof is trivial!*}
+text\<open>Discussed in Avron, Gentzen-Type Systems, Resolution and Tableaux,
+ JAR 10 (265-281), 1993. Proof is trivial!\<close>
lemma "~((\<exists>x.~P(x)) & ((\<exists>x. P(x)) | (\<exists>x. P(x) & Q(x))) & ~ (\<exists>x. P(x)))"
by blast
-subsection{*Problems requiring quantifier duplication*}
+subsection\<open>Problems requiring quantifier duplication\<close>
-text{*Theorem B of Peter Andrews, Theorem Proving via General Matings,
- JACM 28 (1981).*}
+text\<open>Theorem B of Peter Andrews, Theorem Proving via General Matings,
+ JACM 28 (1981).\<close>
lemma "(\<exists>x. \<forall>y. P(x) <-> P(y)) --> ((\<exists>x. P(x)) <-> (\<forall>y. P(y)))"
by blast
-text{*Needs multiple instantiation of ALL.*}
+text\<open>Needs multiple instantiation of ALL.\<close>
lemma "(\<forall>x. P(x)-->P(f(x))) & P(d)-->P(f(f(f(d))))"
by blast
-text{*Needs double instantiation of the quantifier*}
+text\<open>Needs double instantiation of the quantifier\<close>
lemma "\<exists>x. P(x) --> P(a) & P(b)"
by blast
@@ -139,7 +139,7 @@
lemma "\<exists>x. (\<exists>y. P(y)) --> P(x)"
by blast
-text{*V. Lifschitz, What Is the Inverse Method?, JAR 5 (1989), 1--23. NOT PROVED*}
+text\<open>V. Lifschitz, What Is the Inverse Method?, JAR 5 (1989), 1--23. NOT PROVED\<close>
lemma "\<exists>x x'. \<forall>y. \<exists>z z'.
(~P(y,y) | P(x,x) | ~S(z,x)) &
(S(x,y) | ~S(y,z) | Q(z',z')) &
@@ -148,40 +148,40 @@
-subsection{*Hard examples with quantifiers*}
+subsection\<open>Hard examples with quantifiers\<close>
-text{*18*}
+text\<open>18\<close>
lemma "\<exists>y. \<forall>x. P(y)-->P(x)"
by blast
-text{*19*}
+text\<open>19\<close>
lemma "\<exists>x. \<forall>y z. (P(y)-->Q(z)) --> (P(x)-->Q(x))"
by blast
-text{*20*}
+text\<open>20\<close>
lemma "(\<forall>x y. \<exists>z. \<forall>w. (P(x)&Q(y)-->R(z)&S(w)))
--> (\<exists>x y. P(x) & Q(y)) --> (\<exists>z. R(z))"
by blast
-text{*21*}
+text\<open>21\<close>
lemma "(\<exists>x. P-->Q(x)) & (\<exists>x. Q(x)-->P) --> (\<exists>x. P<->Q(x))"
by blast
-text{*22*}
+text\<open>22\<close>
lemma "(\<forall>x. P <-> Q(x)) --> (P <-> (\<forall>x. Q(x)))"
by blast
-text{*23*}
+text\<open>23\<close>
lemma "(\<forall>x. P | Q(x)) <-> (P | (\<forall>x. Q(x)))"
by blast
-text{*24*}
+text\<open>24\<close>
lemma "~(\<exists>x. S(x)&Q(x)) & (\<forall>x. P(x) --> Q(x)|R(x)) &
(~(\<exists>x. P(x)) --> (\<exists>x. Q(x))) & (\<forall>x. Q(x)|R(x) --> S(x))
--> (\<exists>x. P(x)&R(x))"
by blast
-text{*25*}
+text\<open>25\<close>
lemma "(\<exists>x. P(x)) &
(\<forall>x. L(x) --> ~ (M(x) & R(x))) &
(\<forall>x. P(x) --> (M(x) & L(x))) &
@@ -189,13 +189,13 @@
--> (\<exists>x. Q(x)&P(x))"
by blast
-text{*26*}
+text\<open>26\<close>
lemma "((\<exists>x. p(x)) <-> (\<exists>x. q(x))) &
(\<forall>x. \<forall>y. p(x) & q(y) --> (r(x) <-> s(y)))
--> ((\<forall>x. p(x)-->r(x)) <-> (\<forall>x. q(x)-->s(x)))"
by blast
-text{*27*}
+text\<open>27\<close>
lemma "(\<exists>x. P(x) & ~Q(x)) &
(\<forall>x. P(x) --> R(x)) &
(\<forall>x. M(x) & L(x) --> P(x)) &
@@ -203,63 +203,63 @@
--> (\<forall>x. M(x) --> ~L(x))"
by blast
-text{*28. AMENDED*}
+text\<open>28. AMENDED\<close>
lemma "(\<forall>x. P(x) --> (\<forall>x. Q(x))) &
((\<forall>x. Q(x)|R(x)) --> (\<exists>x. Q(x)&S(x))) &
((\<exists>x. S(x)) --> (\<forall>x. L(x) --> M(x)))
--> (\<forall>x. P(x) & L(x) --> M(x))"
by blast
-text{*29. Essentially the same as Principia Mathematica *11.71*}
+text\<open>29. Essentially the same as Principia Mathematica *11.71\<close>
lemma "(\<exists>x. P(x)) & (\<exists>y. Q(y))
--> ((\<forall>x. P(x)-->R(x)) & (\<forall>y. Q(y)-->S(y)) <->
(\<forall>x y. P(x) & Q(y) --> R(x) & S(y)))"
by blast
-text{*30*}
+text\<open>30\<close>
lemma "(\<forall>x. P(x) | Q(x) --> ~ R(x)) &
(\<forall>x. (Q(x) --> ~ S(x)) --> P(x) & R(x))
--> (\<forall>x. S(x))"
by blast
-text{*31*}
+text\<open>31\<close>
lemma "~(\<exists>x. P(x) & (Q(x) | R(x))) &
(\<exists>x. L(x) & P(x)) &
(\<forall>x. ~ R(x) --> M(x))
--> (\<exists>x. L(x) & M(x))"
by blast
-text{*32*}
+text\<open>32\<close>
lemma "(\<forall>x. P(x) & (Q(x)|R(x))-->S(x)) &
(\<forall>x. S(x) & R(x) --> L(x)) &
(\<forall>x. M(x) --> R(x))
--> (\<forall>x. P(x) & M(x) --> L(x))"
by blast
-text{*33*}
+text\<open>33\<close>
lemma "(\<forall>x. P(a) & (P(x)-->P(b))-->P(c)) <->
(\<forall>x. (~P(a) | P(x) | P(c)) & (~P(a) | ~P(b) | P(c)))"
by blast
-text{*34 AMENDED (TWICE!!). Andrews's challenge*}
+text\<open>34 AMENDED (TWICE!!). Andrews's challenge\<close>
lemma "((\<exists>x. \<forall>y. p(x) <-> p(y)) <->
((\<exists>x. q(x)) <-> (\<forall>y. p(y)))) <->
((\<exists>x. \<forall>y. q(x) <-> q(y)) <->
((\<exists>x. p(x)) <-> (\<forall>y. q(y))))"
by blast
-text{*35*}
+text\<open>35\<close>
lemma "\<exists>x y. P(x,y) --> (\<forall>u v. P(u,v))"
by blast
-text{*36*}
+text\<open>36\<close>
lemma "(\<forall>x. \<exists>y. J(x,y)) &
(\<forall>x. \<exists>y. G(x,y)) &
(\<forall>x y. J(x,y) | G(x,y) --> (\<forall>z. J(y,z) | G(y,z) --> H(x,z)))
--> (\<forall>x. \<exists>y. H(x,y))"
by blast
-text{*37*}
+text\<open>37\<close>
lemma "(\<forall>z. \<exists>w. \<forall>x. \<exists>y.
(P(x,z)-->P(y,w)) & P(y,z) & (P(y,w) --> (\<exists>u. Q(u,w)))) &
(\<forall>x z. ~P(x,z) --> (\<exists>y. Q(y,z))) &
@@ -267,7 +267,7 @@
--> (\<forall>x. \<exists>y. R(x,y))"
by blast
-text{*38*}
+text\<open>38\<close>
lemma "(\<forall>x. p(a) & (p(x) --> (\<exists>y. p(y) & r(x,y))) -->
(\<exists>z. \<exists>w. p(z) & r(x,w) & r(w,z))) <->
(\<forall>x. (~p(a) | p(x) | (\<exists>z. \<exists>w. p(z) & r(x,w) & r(w,z))) &
@@ -275,25 +275,25 @@
(\<exists>z. \<exists>w. p(z) & r(x,w) & r(w,z))))"
by blast
-text{*39*}
+text\<open>39\<close>
lemma "~ (\<exists>x. \<forall>y. F(y,x) <-> ~F(y,y))"
by blast
-text{*40. AMENDED*}
+text\<open>40. AMENDED\<close>
lemma "(\<exists>y. \<forall>x. F(x,y) <-> F(x,x)) -->
~(\<forall>x. \<exists>y. \<forall>z. F(z,y) <-> ~ F(z,x))"
by blast
-text{*41*}
+text\<open>41\<close>
lemma "(\<forall>z. \<exists>y. \<forall>x. f(x,y) <-> f(x,z) & ~ f(x,x))
--> ~ (\<exists>z. \<forall>x. f(x,z))"
by blast
-text{*42*}
+text\<open>42\<close>
lemma "~ (\<exists>y. \<forall>x. p(x,y) <-> ~ (\<exists>z. p(x,z) & p(z,x)))"
by blast
-text{*43*}
+text\<open>43\<close>
lemma "(\<forall>x. \<forall>y. q(x,y) <-> (\<forall>z. p(z,x) <-> p(z,y)))
--> (\<forall>x. \<forall>y. q(x,y) <-> q(y,x))"
by blast
@@ -302,13 +302,13 @@
Deepen_tac alone requires 253 secs. Or
by (mini_tac @{context} 1 THEN Deepen_tac 5 1) *)
-text{*44*}
+text\<open>44\<close>
lemma "(\<forall>x. f(x) --> (\<exists>y. g(y) & h(x,y) & (\<exists>y. g(y) & ~ h(x,y)))) &
(\<exists>x. j(x) & (\<forall>y. g(y) --> h(x,y)))
--> (\<exists>x. j(x) & ~f(x))"
by blast
-text{*45*}
+text\<open>45\<close>
lemma "(\<forall>x. f(x) & (\<forall>y. g(y) & h(x,y) --> j(x,y))
--> (\<forall>y. g(y) & h(x,y) --> k(y))) &
~ (\<exists>y. l(y) & k(y)) &
@@ -318,7 +318,7 @@
by blast
-text{*46*}
+text\<open>46\<close>
lemma "(\<forall>x. f(x) & (\<forall>y. f(y) & h(y,x) --> g(y)) --> g(x)) &
((\<exists>x. f(x) & ~g(x)) -->
(\<exists>x. f(x) & ~g(x) & (\<forall>y. f(y) & ~g(y) --> j(x,y)))) &
@@ -327,42 +327,42 @@
by blast
-subsection{*Problems (mainly) involving equality or functions*}
+subsection\<open>Problems (mainly) involving equality or functions\<close>
-text{*48*}
+text\<open>48\<close>
lemma "(a=b | c=d) & (a=c | b=d) --> a=d | b=c"
by blast
-text{*49 NOT PROVED AUTOMATICALLY. Hard because it involves substitution
+text\<open>49 NOT PROVED AUTOMATICALLY. Hard because it involves substitution
for Vars
- the type constraint ensures that x,y,z have the same type as a,b,u. *}
+ the type constraint ensures that x,y,z have the same type as a,b,u.\<close>
lemma "(\<exists>x y::'a. \<forall>z. z=x | z=y) & P(a) & P(b) & a~=b
--> (\<forall>u::'a. P(u))"
apply safe
apply (rule_tac x = a in allE, assumption)
apply (rule_tac x = b in allE, assumption, fast)
- --{*blast's treatment of equality can't do it*}
+ --\<open>blast's treatment of equality can't do it\<close>
done
-text{*50. (What has this to do with equality?) *}
+text\<open>50. (What has this to do with equality?)\<close>
lemma "(\<forall>x. P(a,x) | (\<forall>y. P(x,y))) --> (\<exists>x. \<forall>y. P(x,y))"
by blast
-text{*51*}
+text\<open>51\<close>
lemma "(\<exists>z w. \<forall>x y. P(x,y) <-> (x=z & y=w)) -->
(\<exists>z. \<forall>x. \<exists>w. (\<forall>y. P(x,y) <-> y=w) <-> x=z)"
by blast
-text{*52*}
-text{*Almost the same as 51. *}
+text\<open>52\<close>
+text\<open>Almost the same as 51.\<close>
lemma "(\<exists>z w. \<forall>x y. P(x,y) <-> (x=z & y=w)) -->
(\<exists>w. \<forall>y. \<exists>z. (\<forall>x. P(x,y) <-> x=z) <-> y=w)"
by blast
-text{*55*}
+text\<open>55\<close>
-text{*Non-equational version, from Manthey and Bry, CADE-9 (Springer, 1988).
- fast DISCOVERS who killed Agatha. *}
+text\<open>Non-equational version, from Manthey and Bry, CADE-9 (Springer, 1988).
+ fast DISCOVERS who killed Agatha.\<close>
schematic_lemma "lives(agatha) & lives(butler) & lives(charles) &
(killed(agatha,agatha) | killed(butler,agatha) | killed(charles,agatha)) &
(\<forall>x y. killed(x,y) --> hates(x,y) & ~richer(x,y)) &
@@ -372,53 +372,53 @@
(\<forall>x. hates(agatha,x) --> hates(butler,x)) &
(\<forall>x. ~hates(x,agatha) | ~hates(x,butler) | ~hates(x,charles)) -->
killed(?who,agatha)"
-by fast --{*MUCH faster than blast*}
+by fast --\<open>MUCH faster than blast\<close>
-text{*56*}
+text\<open>56\<close>
lemma "(\<forall>x. (\<exists>y. P(y) & x=f(y)) --> P(x)) <-> (\<forall>x. P(x) --> P(f(x)))"
by blast
-text{*57*}
+text\<open>57\<close>
lemma "P(f(a,b), f(b,c)) & P(f(b,c), f(a,c)) &
(\<forall>x y z. P(x,y) & P(y,z) --> P(x,z)) --> P(f(a,b), f(a,c))"
by blast
-text{*58 NOT PROVED AUTOMATICALLY*}
+text\<open>58 NOT PROVED AUTOMATICALLY\<close>
lemma "(\<forall>x y. f(x)=g(y)) --> (\<forall>x y. f(f(x))=f(g(y)))"
by (slow elim: subst_context)
-text{*59*}
+text\<open>59\<close>
lemma "(\<forall>x. P(x) <-> ~P(f(x))) --> (\<exists>x. P(x) & ~P(f(x)))"
by blast
-text{*60*}
+text\<open>60\<close>
lemma "\<forall>x. P(x,f(x)) <-> (\<exists>y. (\<forall>z. P(z,y) --> P(z,f(x))) & P(x,y))"
by blast
-text{*62 as corrected in JAR 18 (1997), page 135*}
+text\<open>62 as corrected in JAR 18 (1997), page 135\<close>
lemma "(\<forall>x. p(a) & (p(x) --> p(f(x))) --> p(f(f(x)))) <->
(\<forall>x. (~p(a) | p(x) | p(f(f(x)))) &
(~p(a) | ~p(f(x)) | p(f(f(x)))))"
by blast
-text{*From Davis, Obvious Logical Inferences, IJCAI-81, 530-531
- fast indeed copes!*}
+text\<open>From Davis, Obvious Logical Inferences, IJCAI-81, 530-531
+ fast indeed copes!\<close>
lemma "(\<forall>x. F(x) & ~G(x) --> (\<exists>y. H(x,y) & J(y))) &
(\<exists>x. K(x) & F(x) & (\<forall>y. H(x,y) --> K(y))) &
(\<forall>x. K(x) --> ~G(x)) --> (\<exists>x. K(x) & J(x))"
by fast
-text{*From Rudnicki, Obvious Inferences, JAR 3 (1987), 383-393.
- It does seem obvious!*}
+text\<open>From Rudnicki, Obvious Inferences, JAR 3 (1987), 383-393.
+ It does seem obvious!\<close>
lemma "(\<forall>x. F(x) & ~G(x) --> (\<exists>y. H(x,y) & J(y))) &
(\<exists>x. K(x) & F(x) & (\<forall>y. H(x,y) --> K(y))) &
(\<forall>x. K(x) --> ~G(x)) --> (\<exists>x. K(x) --> ~G(x))"
by fast
-text{*Halting problem: Formulation of Li Dafa (AAR Newsletter 27, Oct 1994.)
- author U. Egly*}
+text\<open>Halting problem: Formulation of Li Dafa (AAR Newsletter 27, Oct 1994.)
+ author U. Egly\<close>
lemma "((\<exists>x. A(x) & (\<forall>y. C(y) --> (\<forall>z. D(x,y,z)))) -->
(\<exists>w. C(w) & (\<forall>y. C(y) --> (\<forall>z. D(w,y,z)))))
&
@@ -437,10 +437,10 @@
-->
~ (\<exists>x. A(x) & (\<forall>y. C(y) --> (\<forall>z. D(x,y,z))))"
by (blast 12)
- --{*Needed because the search for depths below 12 is very slow*}
+ --\<open>Needed because the search for depths below 12 is very slow\<close>
-text{*Halting problem II: credited to M. Bruschi by Li Dafa in JAR 18(1), p.105*}
+text\<open>Halting problem II: credited to M. Bruschi by Li Dafa in JAR 18(1), p.105\<close>
lemma "((\<exists>x. A(x) & (\<forall>y. C(y) --> (\<forall>z. D(x,y,z)))) -->
(\<exists>w. C(w) & (\<forall>y. C(y) --> (\<forall>z. D(w,y,z)))))
&
@@ -464,12 +464,12 @@
~ (\<exists>x. A(x) & (\<forall>y. C(y) --> (\<forall>z. D(x,y,z))))"
by blast
-text{* Challenge found on info-hol *}
+text\<open>Challenge found on info-hol\<close>
lemma "\<forall>x. \<exists>v w. \<forall>y z. P(x) & Q(y) --> (P(v) | R(w)) & (R(z) --> Q(v))"
by blast
-text{*Attributed to Lewis Carroll by S. G. Pulman. The first or last assumption
-can be deleted.*}
+text\<open>Attributed to Lewis Carroll by S. G. Pulman. The first or last assumption
+can be deleted.\<close>
lemma "(\<forall>x. honest(x) & industrious(x) --> healthy(x)) &
~ (\<exists>x. grocer(x) & healthy(x)) &
(\<forall>x. industrious(x) & grocer(x) --> honest(x)) &
--- a/src/FOL/ex/First_Order_Logic.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/FOL/ex/First_Order_Logic.thy Thu Jul 23 14:25:05 2015 +0200
@@ -2,18 +2,18 @@
Author: Markus Wenzel, TU Munich
*)
-section {* A simple formulation of First-Order Logic *}
+section \<open>A simple formulation of First-Order Logic\<close>
theory First_Order_Logic imports Pure begin
-text {*
+text \<open>
The subsequent theory development illustrates single-sorted
intuitionistic first-order logic with equality, formulated within
the Pure framework. Actually this is not an example of
Isabelle/FOL, but of Isabelle/Pure.
-*}
+\<close>
-subsection {* Syntax *}
+subsection \<open>Syntax\<close>
typedecl i
typedecl o
@@ -22,7 +22,7 @@
Trueprop :: "o \<Rightarrow> prop" ("_" 5)
-subsection {* Propositional logic *}
+subsection \<open>Propositional logic\<close>
axiomatization
false :: o ("\<bottom>") and
@@ -47,8 +47,8 @@
assumes "A \<and> B"
obtains A and B
proof
- from `A \<and> B` show A by (rule conjD1)
- from `A \<and> B` show B by (rule conjD2)
+ from \<open>A \<and> B\<close> show A by (rule conjD1)
+ from \<open>A \<and> B\<close> show B by (rule conjD2)
qed
definition true :: o ("\<top>")
@@ -101,7 +101,7 @@
qed
-subsection {* Equality *}
+subsection \<open>Equality\<close>
axiomatization
equal :: "i \<Rightarrow> i \<Rightarrow> o" (infixl "=" 50)
@@ -119,7 +119,7 @@
qed
-subsection {* Quantifiers *}
+subsection \<open>Quantifiers\<close>
axiomatization
All :: "(i \<Rightarrow> o) \<Rightarrow> o" (binder "\<forall>" 10) and
--- a/src/FOL/ex/Foundation.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/FOL/ex/Foundation.thy Thu Jul 23 14:25:05 2015 +0200
@@ -18,7 +18,7 @@
apply assumption
done
-text {*A form of conj-elimination*}
+text \<open>A form of conj-elimination\<close>
lemma
assumes "A & B"
and "A ==> B ==> C"
@@ -99,7 +99,7 @@
apply (rule refl)?
oops
-text {* Parallel lifting example. *}
+text \<open>Parallel lifting example.\<close>
lemma "EX u. ALL x. EX v. ALL y. EX w. P(u,x,v,y,w)"
apply (rule exI allI)
apply (rule exI allI)
@@ -121,7 +121,7 @@
apply assumption
done
-text {* A bigger demonstration of quantifiers -- not in the paper. *}
+text \<open>A bigger demonstration of quantifiers -- not in the paper.\<close>
lemma "(EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))"
apply (rule impI)
apply (rule allI)
--- a/src/FOL/ex/If.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/FOL/ex/If.thy Thu Jul 23 14:25:05 2015 +0200
@@ -3,7 +3,7 @@
Copyright 1991 University of Cambridge
*)
-section {* First-Order Logic: the 'if' example *}
+section \<open>First-Order Logic: the 'if' example\<close>
theory If imports FOL begin
@@ -28,7 +28,7 @@
apply (rule ifI)
oops
-text{*Trying again from the beginning in order to use @{text blast}*}
+text\<open>Trying again from the beginning in order to use @{text blast}\<close>
declare ifI [intro!]
declare ifE [elim!]
@@ -39,22 +39,22 @@
lemma "if(if(P,Q,R), A, B) <-> if(P, if(Q,A,B), if(R,A,B))"
by blast
-text{*Trying again from the beginning in order to prove from the definitions*}
+text\<open>Trying again from the beginning in order to prove from the definitions\<close>
lemma "if(if(P,Q,R), A, B) <-> if(P, if(Q,A,B), if(R,A,B))"
by (simp add: if_def, blast)
-text{*An invalid formula. High-level rules permit a simpler diagnosis*}
+text\<open>An invalid formula. High-level rules permit a simpler diagnosis\<close>
lemma "if(if(P,Q,R), A, B) <-> if(P, if(Q,A,B), if(R,B,A))"
apply auto
- -- {*The next step will fail unless subgoals remain*}
+ -- \<open>The next step will fail unless subgoals remain\<close>
apply (tactic all_tac)
oops
-text{*Trying again from the beginning in order to prove from the definitions*}
+text\<open>Trying again from the beginning in order to prove from the definitions\<close>
lemma "if(if(P,Q,R), A, B) <-> if(P, if(Q,A,B), if(R,B,A))"
apply (simp add: if_def, auto)
- -- {*The next step will fail unless subgoals remain*}
+ -- \<open>The next step will fail unless subgoals remain\<close>
apply (tactic all_tac)
oops
--- a/src/FOL/ex/Intro.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/FOL/ex/Intro.thy Thu Jul 23 14:25:05 2015 +0200
@@ -5,13 +5,13 @@
Derives some inference rules, illustrating the use of definitions.
*)
-section {* Examples for the manual ``Introduction to Isabelle'' *}
+section \<open>Examples for the manual ``Introduction to Isabelle''\<close>
theory Intro
imports FOL
begin
-subsubsection {* Some simple backward proofs *}
+subsubsection \<open>Some simple backward proofs\<close>
lemma mythm: "P|P --> P"
apply (rule impI)
@@ -41,7 +41,7 @@
done
-subsubsection {* Demonstration of @{text "fast"} *}
+subsubsection \<open>Demonstration of @{text "fast"}\<close>
lemma "(EX y. ALL x. J(y,x) <-> ~J(x,x))
--> ~ (ALL x. EX y. ALL z. J(z,y) <-> ~ J(z,x))"
@@ -55,7 +55,7 @@
done
-subsubsection {* Derivation of conjunction elimination rule *}
+subsubsection \<open>Derivation of conjunction elimination rule\<close>
lemma
assumes major: "P&Q"
@@ -67,9 +67,9 @@
done
-subsection {* Derived rules involving definitions *}
+subsection \<open>Derived rules involving definitions\<close>
-text {* Derivation of negation introduction *}
+text \<open>Derivation of negation introduction\<close>
lemma
assumes "P ==> False"
@@ -90,7 +90,7 @@
apply (rule minor)
done
-text {* Alternative proof of the result above *}
+text \<open>Alternative proof of the result above\<close>
lemma
assumes major: "~P"
and minor: P
--- a/src/FOL/ex/Intuitionistic.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/FOL/ex/Intuitionistic.thy Thu Jul 23 14:25:05 2015 +0200
@@ -3,7 +3,7 @@
Copyright 1991 University of Cambridge
*)
-section {* Intuitionistic First-Order Logic *}
+section \<open>Intuitionistic First-Order Logic\<close>
theory Intuitionistic
imports IFOL
@@ -21,7 +21,7 @@
*)
-text{*Metatheorem (for \emph{propositional} formulae):
+text\<open>Metatheorem (for \emph{propositional} formulae):
$P$ is classically provable iff $\neg\neg P$ is intuitionistically provable.
Therefore $\neg P$ is classically provable iff it is intuitionistically
provable.
@@ -34,78 +34,78 @@
intuitionistically. The latter is intuitionistically equivalent to $\neg\neg
Q\rightarrow\neg\neg P$, hence to $\neg\neg P$, since $\neg\neg Q$ is
intuitionistically provable. Finally, if $P$ is a negation then $\neg\neg P$
-is intuitionstically equivalent to $P$. [Andy Pitts] *}
+is intuitionstically equivalent to $P$. [Andy Pitts]\<close>
lemma "~~(P&Q) <-> ~~P & ~~Q"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
lemma "~~ ((~P --> Q) --> (~P --> ~Q) --> P)"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
-text{*Double-negation does NOT distribute over disjunction*}
+text\<open>Double-negation does NOT distribute over disjunction\<close>
lemma "~~(P-->Q) <-> (~~P --> ~~Q)"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
lemma "~~~P <-> ~P"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
lemma "~~((P --> Q | R) --> (P-->Q) | (P-->R))"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
lemma "(P<->Q) <-> (Q<->P)"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
lemma "((P --> (Q | (Q-->R))) --> R) --> R"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
lemma "(((G-->A) --> J) --> D --> E) --> (((H-->B)-->I)-->C-->J)
--> (A-->H) --> F --> G --> (((C-->B)-->I)-->D)-->(A-->C)
--> (((F-->A)-->B) --> I) --> E"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
-text{*Lemmas for the propositional double-negation translation*}
+text\<open>Lemmas for the propositional double-negation translation\<close>
lemma "P --> ~~P"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
lemma "~~(~~P --> P)"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
lemma "~~P & ~~(P --> Q) --> ~~Q"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
-text{*The following are classically but not constructively valid.
- The attempt to prove them terminates quickly!*}
+text\<open>The following are classically but not constructively valid.
+ The attempt to prove them terminates quickly!\<close>
lemma "((P-->Q) --> P) --> P"
-apply (tactic{*IntPr.fast_tac @{context} 1*} | -)
-apply (rule asm_rl) --{*Checks that subgoals remain: proof failed.*}
+apply (tactic\<open>IntPr.fast_tac @{context} 1\<close> | -)
+apply (rule asm_rl) --\<open>Checks that subgoals remain: proof failed.\<close>
oops
lemma "(P&Q-->R) --> (P-->R) | (Q-->R)"
-apply (tactic{*IntPr.fast_tac @{context} 1*} | -)
-apply (rule asm_rl) --{*Checks that subgoals remain: proof failed.*}
+apply (tactic\<open>IntPr.fast_tac @{context} 1\<close> | -)
+apply (rule asm_rl) --\<open>Checks that subgoals remain: proof failed.\<close>
oops
-subsection{*de Bruijn formulae*}
+subsection\<open>de Bruijn formulae\<close>
-text{*de Bruijn formula with three predicates*}
+text\<open>de Bruijn formula with three predicates\<close>
lemma "((P<->Q) --> P&Q&R) &
((Q<->R) --> P&Q&R) &
((R<->P) --> P&Q&R) --> P&Q&R"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
-text{*de Bruijn formula with five predicates*}
+text\<open>de Bruijn formula with five predicates\<close>
lemma "((P<->Q) --> P&Q&R&S&T) &
((Q<->R) --> P&Q&R&S&T) &
((R<->S) --> P&Q&R&S&T) &
((S<->T) --> P&Q&R&S&T) &
((T<->P) --> P&Q&R&S&T) --> P&Q&R&S&T"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
(*** Problems from of Sahlin, Franzen and Haridi,
@@ -113,313 +113,313 @@
J. Logic and Comp. 2 (5), October 1992, 619-656.
***)
-text{*Problem 1.1*}
+text\<open>Problem 1.1\<close>
lemma "(ALL x. EX y. ALL z. p(x) & q(y) & r(z)) <->
(ALL z. EX y. ALL x. p(x) & q(y) & r(z))"
-by (tactic{*IntPr.best_dup_tac @{context} 1*}) --{*SLOW*}
+by (tactic\<open>IntPr.best_dup_tac @{context} 1\<close>) --\<open>SLOW\<close>
-text{*Problem 3.1*}
+text\<open>Problem 3.1\<close>
lemma "~ (EX x. ALL y. mem(y,x) <-> ~ mem(x,x))"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
-text{*Problem 4.1: hopeless!*}
+text\<open>Problem 4.1: hopeless!\<close>
lemma "(ALL x. p(x) --> p(h(x)) | p(g(x))) & (EX x. p(x)) & (ALL x. ~p(h(x)))
--> (EX x. p(g(g(g(g(g(x)))))))"
oops
-subsection{*Intuitionistic FOL: propositional problems based on Pelletier.*}
+subsection\<open>Intuitionistic FOL: propositional problems based on Pelletier.\<close>
-text{*~~1*}
+text\<open>~~1\<close>
lemma "~~((P-->Q) <-> (~Q --> ~P))"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
-text{*~~2*}
+text\<open>~~2\<close>
lemma "~~(~~P <-> P)"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
-text{*3*}
+text\<open>3\<close>
lemma "~(P-->Q) --> (Q-->P)"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
-text{*~~4*}
+text\<open>~~4\<close>
lemma "~~((~P-->Q) <-> (~Q --> P))"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
-text{*~~5*}
+text\<open>~~5\<close>
lemma "~~((P|Q-->P|R) --> P|(Q-->R))"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
-text{*~~6*}
+text\<open>~~6\<close>
lemma "~~(P | ~P)"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
-text{*~~7*}
+text\<open>~~7\<close>
lemma "~~(P | ~~~P)"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
-text{*~~8. Peirce's law*}
+text\<open>~~8. Peirce's law\<close>
lemma "~~(((P-->Q) --> P) --> P)"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
-text{*9*}
+text\<open>9\<close>
lemma "((P|Q) & (~P|Q) & (P| ~Q)) --> ~ (~P | ~Q)"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
-text{*10*}
+text\<open>10\<close>
lemma "(Q-->R) --> (R-->P&Q) --> (P-->(Q|R)) --> (P<->Q)"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
-subsection{*11. Proved in each direction (incorrectly, says Pelletier!!) *}
+subsection\<open>11. Proved in each direction (incorrectly, says Pelletier!!)\<close>
lemma "P<->P"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
-text{*~~12. Dijkstra's law *}
+text\<open>~~12. Dijkstra's law\<close>
lemma "~~(((P <-> Q) <-> R) <-> (P <-> (Q <-> R)))"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
lemma "((P <-> Q) <-> R) --> ~~(P <-> (Q <-> R))"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
-text{*13. Distributive law*}
+text\<open>13. Distributive law\<close>
lemma "P | (Q & R) <-> (P | Q) & (P | R)"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
-text{*~~14*}
+text\<open>~~14\<close>
lemma "~~((P <-> Q) <-> ((Q | ~P) & (~Q|P)))"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
-text{*~~15*}
+text\<open>~~15\<close>
lemma "~~((P --> Q) <-> (~P | Q))"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
-text{*~~16*}
+text\<open>~~16\<close>
lemma "~~((P-->Q) | (Q-->P))"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
-text{*~~17*}
+text\<open>~~17\<close>
lemma "~~(((P & (Q-->R))-->S) <-> ((~P | Q | S) & (~P | ~R | S)))"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
-text{*Dijkstra's "Golden Rule"*}
+text\<open>Dijkstra's "Golden Rule"\<close>
lemma "(P&Q) <-> P <-> Q <-> (P|Q)"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
-subsection{*****Examples with quantifiers*****}
+subsection\<open>****Examples with quantifiers****\<close>
-subsection{*The converse is classical in the following implications...*}
+subsection\<open>The converse is classical in the following implications...\<close>
lemma "(EX x. P(x)-->Q) --> (ALL x. P(x)) --> Q"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
lemma "((ALL x. P(x))-->Q) --> ~ (ALL x. P(x) & ~Q)"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
lemma "((ALL x. ~P(x))-->Q) --> ~ (ALL x. ~ (P(x)|Q))"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
lemma "(ALL x. P(x)) | Q --> (ALL x. P(x) | Q)"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
lemma "(EX x. P --> Q(x)) --> (P --> (EX x. Q(x)))"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
-subsection{*The following are not constructively valid!*}
-text{*The attempt to prove them terminates quickly!*}
+subsection\<open>The following are not constructively valid!\<close>
+text\<open>The attempt to prove them terminates quickly!\<close>
lemma "((ALL x. P(x))-->Q) --> (EX x. P(x)-->Q)"
-apply (tactic{*IntPr.fast_tac @{context} 1*} | -)
-apply (rule asm_rl) --{*Checks that subgoals remain: proof failed.*}
+apply (tactic\<open>IntPr.fast_tac @{context} 1\<close> | -)
+apply (rule asm_rl) --\<open>Checks that subgoals remain: proof failed.\<close>
oops
lemma "(P --> (EX x. Q(x))) --> (EX x. P-->Q(x))"
-apply (tactic{*IntPr.fast_tac @{context} 1*} | -)
-apply (rule asm_rl) --{*Checks that subgoals remain: proof failed.*}
+apply (tactic\<open>IntPr.fast_tac @{context} 1\<close> | -)
+apply (rule asm_rl) --\<open>Checks that subgoals remain: proof failed.\<close>
oops
lemma "(ALL x. P(x) | Q) --> ((ALL x. P(x)) | Q)"
-apply (tactic{*IntPr.fast_tac @{context} 1*} | -)
-apply (rule asm_rl) --{*Checks that subgoals remain: proof failed.*}
+apply (tactic\<open>IntPr.fast_tac @{context} 1\<close> | -)
+apply (rule asm_rl) --\<open>Checks that subgoals remain: proof failed.\<close>
oops
lemma "(ALL x. ~~P(x)) --> ~~(ALL x. P(x))"
-apply (tactic{*IntPr.fast_tac @{context} 1*} | -)
-apply (rule asm_rl) --{*Checks that subgoals remain: proof failed.*}
+apply (tactic\<open>IntPr.fast_tac @{context} 1\<close> | -)
+apply (rule asm_rl) --\<open>Checks that subgoals remain: proof failed.\<close>
oops
-text{*Classically but not intuitionistically valid. Proved by a bug in 1986!*}
+text\<open>Classically but not intuitionistically valid. Proved by a bug in 1986!\<close>
lemma "EX x. Q(x) --> (ALL x. Q(x))"
-apply (tactic{*IntPr.fast_tac @{context} 1*} | -)
-apply (rule asm_rl) --{*Checks that subgoals remain: proof failed.*}
+apply (tactic\<open>IntPr.fast_tac @{context} 1\<close> | -)
+apply (rule asm_rl) --\<open>Checks that subgoals remain: proof failed.\<close>
oops
-subsection{*Hard examples with quantifiers*}
+subsection\<open>Hard examples with quantifiers\<close>
-text{*The ones that have not been proved are not known to be valid!
- Some will require quantifier duplication -- not currently available*}
+text\<open>The ones that have not been proved are not known to be valid!
+ Some will require quantifier duplication -- not currently available\<close>
-text{*~~18*}
+text\<open>~~18\<close>
lemma "~~(EX y. ALL x. P(y)-->P(x))"
-oops --{*NOT PROVED*}
+oops --\<open>NOT PROVED\<close>
-text{*~~19*}
+text\<open>~~19\<close>
lemma "~~(EX x. ALL y z. (P(y)-->Q(z)) --> (P(x)-->Q(x)))"
-oops --{*NOT PROVED*}
+oops --\<open>NOT PROVED\<close>
-text{*20*}
+text\<open>20\<close>
lemma "(ALL x y. EX z. ALL w. (P(x)&Q(y)-->R(z)&S(w)))
--> (EX x y. P(x) & Q(y)) --> (EX z. R(z))"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
-text{*21*}
+text\<open>21\<close>
lemma "(EX x. P-->Q(x)) & (EX x. Q(x)-->P) --> ~~(EX x. P<->Q(x))"
-oops --{*NOT PROVED; needs quantifier duplication*}
+oops --\<open>NOT PROVED; needs quantifier duplication\<close>
-text{*22*}
+text\<open>22\<close>
lemma "(ALL x. P <-> Q(x)) --> (P <-> (ALL x. Q(x)))"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
-text{*~~23*}
+text\<open>~~23\<close>
lemma "~~ ((ALL x. P | Q(x)) <-> (P | (ALL x. Q(x))))"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
-text{*24*}
+text\<open>24\<close>
lemma "~(EX x. S(x)&Q(x)) & (ALL x. P(x) --> Q(x)|R(x)) &
(~(EX x. P(x)) --> (EX x. Q(x))) & (ALL x. Q(x)|R(x) --> S(x))
--> ~~(EX x. P(x)&R(x))"
-txt{*Not clear why @{text fast_tac}, @{text best_tac}, @{text ASTAR} and
- @{text ITER_DEEPEN} all take forever*}
-apply (tactic{* IntPr.safe_tac @{context}*})
+txt\<open>Not clear why @{text fast_tac}, @{text best_tac}, @{text ASTAR} and
+ @{text ITER_DEEPEN} all take forever\<close>
+apply (tactic\<open>IntPr.safe_tac @{context}\<close>)
apply (erule impE)
-apply (tactic{*IntPr.fast_tac @{context} 1*})
-by (tactic{*IntPr.fast_tac @{context} 1*})
+apply (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
-text{*25*}
+text\<open>25\<close>
lemma "(EX x. P(x)) &
(ALL x. L(x) --> ~ (M(x) & R(x))) &
(ALL x. P(x) --> (M(x) & L(x))) &
((ALL x. P(x)-->Q(x)) | (EX x. P(x)&R(x)))
--> (EX x. Q(x)&P(x))"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
-text{*~~26*}
+text\<open>~~26\<close>
lemma "(~~(EX x. p(x)) <-> ~~(EX x. q(x))) &
(ALL x. ALL y. p(x) & q(y) --> (r(x) <-> s(y)))
--> ((ALL x. p(x)-->r(x)) <-> (ALL x. q(x)-->s(x)))"
-oops --{*NOT PROVED*}
+oops --\<open>NOT PROVED\<close>
-text{*27*}
+text\<open>27\<close>
lemma "(EX x. P(x) & ~Q(x)) &
(ALL x. P(x) --> R(x)) &
(ALL x. M(x) & L(x) --> P(x)) &
((EX x. R(x) & ~ Q(x)) --> (ALL x. L(x) --> ~ R(x)))
--> (ALL x. M(x) --> ~L(x))"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
-text{*~~28. AMENDED*}
+text\<open>~~28. AMENDED\<close>
lemma "(ALL x. P(x) --> (ALL x. Q(x))) &
(~~(ALL x. Q(x)|R(x)) --> (EX x. Q(x)&S(x))) &
(~~(EX x. S(x)) --> (ALL x. L(x) --> M(x)))
--> (ALL x. P(x) & L(x) --> M(x))"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
-text{*29. Essentially the same as Principia Mathematica *11.71*}
+text\<open>29. Essentially the same as Principia Mathematica *11.71\<close>
lemma "(EX x. P(x)) & (EX y. Q(y))
--> ((ALL x. P(x)-->R(x)) & (ALL y. Q(y)-->S(y)) <->
(ALL x y. P(x) & Q(y) --> R(x) & S(y)))"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
-text{*~~30*}
+text\<open>~~30\<close>
lemma "(ALL x. (P(x) | Q(x)) --> ~ R(x)) &
(ALL x. (Q(x) --> ~ S(x)) --> P(x) & R(x))
--> (ALL x. ~~S(x))"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
-text{*31*}
+text\<open>31\<close>
lemma "~(EX x. P(x) & (Q(x) | R(x))) &
(EX x. L(x) & P(x)) &
(ALL x. ~ R(x) --> M(x))
--> (EX x. L(x) & M(x))"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
-text{*32*}
+text\<open>32\<close>
lemma "(ALL x. P(x) & (Q(x)|R(x))-->S(x)) &
(ALL x. S(x) & R(x) --> L(x)) &
(ALL x. M(x) --> R(x))
--> (ALL x. P(x) & M(x) --> L(x))"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
-text{*~~33*}
+text\<open>~~33\<close>
lemma "(ALL x. ~~(P(a) & (P(x)-->P(b))-->P(c))) <->
(ALL x. ~~((~P(a) | P(x) | P(c)) & (~P(a) | ~P(b) | P(c))))"
-apply (tactic{*IntPr.best_tac @{context} 1*})
+apply (tactic\<open>IntPr.best_tac @{context} 1\<close>)
done
-text{*36*}
+text\<open>36\<close>
lemma "(ALL x. EX y. J(x,y)) &
(ALL x. EX y. G(x,y)) &
(ALL x y. J(x,y) | G(x,y) --> (ALL z. J(y,z) | G(y,z) --> H(x,z)))
--> (ALL x. EX y. H(x,y))"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
-text{*37*}
+text\<open>37\<close>
lemma "(ALL z. EX w. ALL x. EX y.
~~(P(x,z)-->P(y,w)) & P(y,z) & (P(y,w) --> (EX u. Q(u,w)))) &
(ALL x z. ~P(x,z) --> (EX y. Q(y,z))) &
(~~(EX x y. Q(x,y)) --> (ALL x. R(x,x)))
--> ~~(ALL x. EX y. R(x,y))"
-oops --{*NOT PROVED*}
+oops --\<open>NOT PROVED\<close>
-text{*39*}
+text\<open>39\<close>
lemma "~ (EX x. ALL y. F(y,x) <-> ~F(y,y))"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
-text{*40. AMENDED*}
+text\<open>40. AMENDED\<close>
lemma "(EX y. ALL x. F(x,y) <-> F(x,x)) -->
~(ALL x. EX y. ALL z. F(z,y) <-> ~ F(z,x))"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
-text{*44*}
+text\<open>44\<close>
lemma "(ALL x. f(x) -->
(EX y. g(y) & h(x,y) & (EX y. g(y) & ~ h(x,y)))) &
(EX x. j(x) & (ALL y. g(y) --> h(x,y)))
--> (EX x. j(x) & ~f(x))"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
-text{*48*}
+text\<open>48\<close>
lemma "(a=b | c=d) & (a=c | b=d) --> a=d | b=c"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
-text{*51*}
+text\<open>51\<close>
lemma "(EX z w. ALL x y. P(x,y) <-> (x=z & y=w)) -->
(EX z. ALL x. EX w. (ALL y. P(x,y) <-> y=w) <-> x=z)"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
-text{*52*}
-text{*Almost the same as 51. *}
+text\<open>52\<close>
+text\<open>Almost the same as 51.\<close>
lemma "(EX z w. ALL x y. P(x,y) <-> (x=z & y=w)) -->
(EX w. ALL y. EX z. (ALL x. P(x,y) <-> x=z) <-> y=w)"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
-text{*56*}
+text\<open>56\<close>
lemma "(ALL x. (EX y. P(y) & x=f(y)) --> P(x)) <-> (ALL x. P(x) --> P(f(x)))"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
-text{*57*}
+text\<open>57\<close>
lemma "P(f(a,b), f(b,c)) & P(f(b,c), f(a,c)) &
(ALL x y z. P(x,y) & P(y,z) --> P(x,z)) --> P(f(a,b), f(a,c))"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
-text{*60*}
+text\<open>60\<close>
lemma "ALL x. P(x,f(x)) <-> (EX y. (ALL z. P(z,y) --> P(z,f(x))) & P(x,y))"
-by (tactic{*IntPr.fast_tac @{context} 1*})
+by (tactic\<open>IntPr.fast_tac @{context} 1\<close>)
end
--- a/src/FOL/ex/Locale_Test/Locale_Test.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/FOL/ex/Locale_Test/Locale_Test.thy Thu Jul 23 14:25:05 2015 +0200
@@ -8,7 +8,7 @@
imports Locale_Test1 Locale_Test2 Locale_Test3
begin
-text {* Result of theory merge with distinct but identical interpretations *}
+text \<open>Result of theory merge with distinct but identical interpretations\<close>
context mixin_thy_merge
begin
--- a/src/FOL/ex/Locale_Test/Locale_Test1.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/FOL/ex/Locale_Test/Locale_Test1.thy Thu Jul 23 14:25:05 2015 +0200
@@ -21,7 +21,7 @@
int_minus: "(-x) + x = 0" and
int_minus2: "-(-x) = x"
-section {* Inference of parameter types *}
+section \<open>Inference of parameter types\<close>
locale param1 = fixes p
print_locale! param1
@@ -41,7 +41,7 @@
print_locale! param4
-subsection {* Incremental type constraints *}
+subsection \<open>Incremental type constraints\<close>
locale constraint1 =
fixes prod (infixl "**" 65)
@@ -55,7 +55,7 @@
print_locale! constraint2
-section {* Inheritance *}
+section \<open>Inheritance\<close>
locale semi =
fixes prod (infixl "**" 65)
@@ -86,12 +86,12 @@
locale pert_hom = d1: perturbation f d1 + d2: perturbation f d2 for f d1 d2
print_locale! pert_hom thm pert_hom_def
-text {* Alternative expression, obtaining nicer names in @{text "semi f"}. *}
+text \<open>Alternative expression, obtaining nicer names in @{text "semi f"}.\<close>
locale pert_hom' = semi f + d1: perturbation f d1 + d2: perturbation f d2 for f d1 d2
print_locale! pert_hom' thm pert_hom'_def
-section {* Syntax declarations *}
+section \<open>Syntax declarations\<close>
locale logic =
fixes land (infixl "&&" 55)
@@ -126,7 +126,7 @@
thm var.test_def
-text {* Under which circumstances term syntax remains active. *}
+text \<open>Under which circumstances term syntax remains active.\<close>
locale "syntax" =
fixes p1 :: "'a => 'b"
@@ -147,7 +147,7 @@
thm d1_def d2_def (* should print as "d1(?x) <-> ..." and "d2(?x) <-> ..." *)
-ML {*
+ML \<open>
fun check_syntax ctxt thm expected =
let
val obtained =
@@ -157,14 +157,14 @@
then error ("Theorem syntax '" ^ obtained ^ "' obtained, but '" ^ expected ^ "' expected.")
else ()
end;
-*}
+\<close>
declare [[show_hyps]]
-ML {*
+ML \<open>
check_syntax @{context} @{thm d1_def} "d1(?x) <-> ~ p2(p1(?x))";
check_syntax @{context} @{thm d2_def} "d2(?x) <-> ~ p2(?x)";
-*}
+\<close>
end
@@ -174,21 +174,21 @@
thm d1_def d2_def
(* should print as "syntax.d1(p3, p2, ?x) <-> ..." and "d2(?x) <-> ..." *)
-ML {*
+ML \<open>
check_syntax @{context} @{thm d1_def} "syntax.d1(p3, p2, ?x) <-> ~ p2(p3(?x))";
check_syntax @{context} @{thm d2_def} "d2(?x) <-> ~ p2(?x)";
-*}
+\<close>
end
-section {* Foundational versions of theorems *}
+section \<open>Foundational versions of theorems\<close>
thm logic.assoc
thm logic.lor_def
-section {* Defines *}
+section \<open>Defines\<close>
locale logic_def =
fixes land (infixl "&&" 55)
@@ -217,7 +217,7 @@
end
-section {* Notes *}
+section \<open>Notes\<close>
(* A somewhat arcane homomorphism example *)
@@ -247,7 +247,7 @@
shows True ..
-section {* Theorem statements *}
+section \<open>Theorem statements\<close>
lemma (in lgrp) lcancel:
"x ** y = x ** z <-> y = z"
@@ -278,13 +278,13 @@
print_locale! rgrp
-subsection {* Patterns *}
+subsection \<open>Patterns\<close>
lemma (in rgrp)
assumes "y ** x = z ** x" (is ?a)
shows "y = z" (is ?t)
proof -
- txt {* Weird proof involving patterns from context element and conclusion. *}
+ txt \<open>Weird proof involving patterns from context element and conclusion.\<close>
{
assume ?a
then have "y ** (x ** inv(x)) = z ** (x ** inv(x))"
@@ -292,11 +292,11 @@
then have ?t by (simp add: rone rinv)
}
note x = this
- show ?t by (rule x [OF `?a`])
+ show ?t by (rule x [OF \<open>?a\<close>])
qed
-section {* Interpretation between locales: sublocales *}
+section \<open>Interpretation between locales: sublocales\<close>
sublocale lgrp < right: rgrp
print_facts
@@ -436,7 +436,7 @@
print_locale! trivial (* No instance for y created (subsumed). *)
-subsection {* Sublocale, then interpretation in theory *}
+subsection \<open>Sublocale, then interpretation in theory\<close>
interpretation int?: lgrp "op +" "0" "minus"
proof unfold_locales
@@ -447,15 +447,15 @@
interpretation int2?: semi "op +"
by unfold_locales (* subsumed, thm int2.assoc not generated *)
-ML {* (Global_Theory.get_thms @{theory} "int2.assoc";
+ML \<open>(Global_Theory.get_thms @{theory} "int2.assoc";
raise Fail "thm int2.assoc was generated")
- handle ERROR _ => ([]:thm list); *}
+ handle ERROR _ => ([]:thm list);\<close>
thm int.lone int.right.rone
(* the latter comes through the sublocale relation *)
-subsection {* Interpretation in theory, then sublocale *}
+subsection \<open>Interpretation in theory, then sublocale\<close>
interpretation fol: logic "op +" "minus"
by unfold_locales (rule int_assoc int_minus2)+
@@ -478,7 +478,7 @@
thm fol.two.assoc
-subsection {* Declarations and sublocale *}
+subsection \<open>Declarations and sublocale\<close>
locale logic_a = logic
locale logic_b = logic
@@ -487,9 +487,9 @@
by unfold_locales
-subsection {* Interpretation *}
+subsection \<open>Interpretation\<close>
-subsection {* Rewrite morphism *}
+subsection \<open>Rewrite morphism\<close>
locale logic_o =
fixes land (infixl "&&" 55)
@@ -521,7 +521,7 @@
x.lor_triv
-subsection {* Inheritance of rewrite morphisms *}
+subsection \<open>Inheritance of rewrite morphisms\<close>
locale reflexive =
fixes le :: "'a => 'a => o" (infix "\<sqsubseteq>" 50)
@@ -539,7 +539,7 @@
grefl: "gle(x, x)" and gless_def: "gless(x, y) <-> gle(x, y) & x ~= y" and
grefl': "gle'(x, x)" and gless'_def: "gless'(x, y) <-> gle'(x, y) & x ~= y"
-text {* Setup *}
+text \<open>Setup\<close>
locale mixin = reflexive
begin
@@ -554,7 +554,7 @@
by (simp add: reflexive.less_def[OF reflexive] gless_def)
qed
-text {* Rewrite morphism propagated along the locale hierarchy *}
+text \<open>Rewrite morphism propagated along the locale hierarchy\<close>
locale mixin2 = mixin
begin
@@ -568,7 +568,7 @@
lemma "gless(x, y) <-> gle(x, y) & x ~= y"
by (rule le.less_thm2)
-text {* Rewrite morphism does not leak to a side branch. *}
+text \<open>Rewrite morphism does not leak to a side branch.\<close>
locale mixin3 = reflexive
begin
@@ -581,7 +581,7 @@
thm le.less_thm3 (* rewrite morphism not applied *)
lemma "reflexive.less(gle, x, y) <-> gle(x, y) & x ~= y" by (rule le.less_thm3)
-text {* Rewrite morphism only available in original context *}
+text \<open>Rewrite morphism only available in original context\<close>
locale mixin4_base = reflexive
@@ -613,7 +613,7 @@
lemma "reflexive.less(gle, x, y) <-> gle(x, y) & x ~= y"
by (rule le4.less_thm4')
-text {* Inherited rewrite morphism applied to new theorem *}
+text \<open>Inherited rewrite morphism applied to new theorem\<close>
locale mixin5_base = reflexive
@@ -637,7 +637,7 @@
lemma "gless(x, y) <-> gle(x, y) & x ~= y"
by (rule le5.less_thm5)
-text {* Rewrite morphism pushed down to existing inherited locale *}
+text \<open>Rewrite morphism pushed down to existing inherited locale\<close>
locale mixin6_base = reflexive
@@ -662,7 +662,7 @@
lemma "gless(x, y) <-> gle(x, y) & x ~= y"
by (rule le6.less_thm6)
-text {* Existing rewrite morphism inherited through sublocale relation *}
+text \<open>Existing rewrite morphism inherited through sublocale relation\<close>
locale mixin7_base = reflexive
@@ -696,16 +696,16 @@
by (rule le7.less_thm7b)
-text {* This locale will be interpreted in later theories. *}
+text \<open>This locale will be interpreted in later theories.\<close>
locale mixin_thy_merge = le: reflexive le + le': reflexive le' for le le'
-subsection {* Rewrite morphisms in sublocale *}
+subsection \<open>Rewrite morphisms in sublocale\<close>
-text {* Simulate a specification of left groups where unit and inverse are defined
+text \<open>Simulate a specification of left groups where unit and inverse are defined
rather than specified. This is possible, but not in FOL, due to the lack of a
- selection operator. *}
+ selection operator.\<close>
axiomatization glob_one and glob_inv
where glob_lone: "prod(glob_one(prod), x) = x"
@@ -745,24 +745,24 @@
context lgrp begin
-text {* Equations stored in target *}
+text \<open>Equations stored in target\<close>
lemma "dgrp.one(prod) = one" by (rule one_equation)
lemma "dgrp.inv(prod, x) = inv(x)" by (rule inv_equation)
-text {* Rewrite morphisms applied *}
+text \<open>Rewrite morphisms applied\<close>
lemma "one = glob_one(prod)" by (rule one_def)
lemma "inv(x) = glob_inv(prod, x)" by (rule inv_def)
end
-text {* Interpreted versions *}
+text \<open>Interpreted versions\<close>
lemma "0 = glob_one (op +)" by (rule int.def.one_def)
lemma "- x = glob_inv(op +, x)" by (rule int.def.inv_def)
-text {* Roundup applies rewrite morphisms at declaration level in DFS tree *}
+text \<open>Roundup applies rewrite morphisms at declaration level in DFS tree\<close>
locale roundup = fixes x assumes true: "x <-> True"
@@ -771,7 +771,7 @@
lemma (in roundup) "True & True <-> True" by (rule sub.true)
-section {* Interpretation in named contexts *}
+section \<open>Interpretation in named contexts\<close>
locale container
begin
@@ -780,15 +780,15 @@
end
context container begin
-ML {* (Context.>> (fn generic => let val context = Context.proof_of generic
+ML \<open>(Context.>> (fn generic => let val context = Context.proof_of generic
val _ = Proof_Context.get_thms context "private.true" in generic end);
raise Fail "thm private.true was persisted")
- handle ERROR _ => ([]:thm list); *}
+ handle ERROR _ => ([]:thm list);\<close>
thm true_copy
end
-section {* Interpretation in proofs *}
+section \<open>Interpretation in proofs\<close>
lemma True
proof
--- a/src/FOL/ex/Miniscope.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/FOL/ex/Miniscope.thy Thu Jul 23 14:25:05 2015 +0200
@@ -14,9 +14,9 @@
lemmas ccontr = FalseE [THEN classical]
-subsection {* Negation Normal Form *}
+subsection \<open>Negation Normal Form\<close>
-subsubsection {* de Morgan laws *}
+subsubsection \<open>de Morgan laws\<close>
lemma demorgans:
"~(P&Q) <-> ~P | ~Q"
@@ -38,7 +38,7 @@
(* BEWARE: rewrite rules for <-> can confuse the simplifier!! *)
-subsubsection {* Pushing in the existential quantifiers *}
+subsubsection \<open>Pushing in the existential quantifiers\<close>
lemma ex_simps:
"(EX x. P) <-> P"
@@ -50,7 +50,7 @@
by blast+
-subsubsection {* Pushing in the universal quantifiers *}
+subsubsection \<open>Pushing in the universal quantifiers\<close>
lemma all_simps:
"(ALL x. P) <-> P"
@@ -63,10 +63,10 @@
lemmas mini_simps = demorgans nnf_simps ex_simps all_simps
-ML {*
+ML \<open>
val mini_ss = simpset_of (@{context} addsimps @{thms mini_simps});
fun mini_tac ctxt =
resolve_tac ctxt @{thms ccontr} THEN' asm_full_simp_tac (put_simpset mini_ss ctxt);
-*}
+\<close>
end
--- a/src/FOL/ex/Nat.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/FOL/ex/Nat.thy Thu Jul 23 14:25:05 2015 +0200
@@ -3,7 +3,7 @@
Copyright 1992 University of Cambridge
*)
-section {* Theory of the natural numbers: Peano's axioms, primitive recursion *}
+section \<open>Theory of the natural numbers: Peano's axioms, primitive recursion\<close>
theory Nat
imports FOL
@@ -27,7 +27,7 @@
where "m + n == rec(m, n, %x y. Suc(y))"
-subsection {* Proofs about the natural numbers *}
+subsection \<open>Proofs about the natural numbers\<close>
lemma Suc_n_not_n: "Suc(k) ~= k"
apply (rule_tac n = k in induct)
--- a/src/FOL/ex/Nat_Class.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/FOL/ex/Nat_Class.thy Thu Jul 23 14:25:05 2015 +0200
@@ -6,14 +6,14 @@
imports FOL
begin
-text {*
+text \<open>
This is an abstract version of theory @{text Nat}. Instead of
axiomatizing a single type @{text nat} we define the class of all
these types (up to isomorphism).
Note: The @{text rec} operator had to be made \emph{monomorphic},
because class axioms may not contain more than one type variable.
-*}
+\<close>
class nat =
fixes Zero :: 'a ("0")
--- a/src/FOL/ex/Natural_Numbers.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/FOL/ex/Natural_Numbers.thy Thu Jul 23 14:25:05 2015 +0200
@@ -2,16 +2,16 @@
Author: Markus Wenzel, TU Munich
*)
-section {* Natural numbers *}
+section \<open>Natural numbers\<close>
theory Natural_Numbers
imports FOL
begin
-text {*
+text \<open>
Theory of the natural numbers: Peano's axioms, primitive recursion.
(Modernized version of Larry Paulson's theory "Nat".) \medskip
-*}
+\<close>
typedecl nat
instance nat :: "term" ..
--- a/src/FOL/ex/Prolog.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/FOL/ex/Prolog.thy Thu Jul 23 14:25:05 2015 +0200
@@ -3,7 +3,7 @@
Copyright 1992 University of Cambridge
*)
-section {* First-Order Logic: PROLOG examples *}
+section \<open>First-Order Logic: PROLOG examples\<close>
theory Prolog
imports FOL
@@ -56,16 +56,16 @@
done
schematic_lemma "rev(?x, a:b:c:Nil)"
-apply (rule rules)+ -- {* does not solve it directly! *}
+apply (rule rules)+ -- \<open>does not solve it directly!\<close>
back
back
done
(*backtracking version*)
-ML {*
+ML \<open>
fun prolog_tac ctxt =
DEPTH_FIRST (has_fewer_prems 1) (resolve_tac ctxt @{thms rules} 1)
-*}
+\<close>
schematic_lemma "rev(?x, a:b:c:Nil)"
apply (tactic \<open>prolog_tac @{context}\<close>)
@@ -77,15 +77,15 @@
(*rev([a..p], ?w) requires 153 inferences *)
schematic_lemma "rev(a:b:c:d:e:f:g:h:i:j:k:l:m:n:o:p:Nil, ?w)"
-apply (tactic {*
- DEPTH_SOLVE (resolve_tac @{context} ([@{thm refl}, @{thm conjI}] @ @{thms rules}) 1) *})
+apply (tactic \<open>
+ DEPTH_SOLVE (resolve_tac @{context} ([@{thm refl}, @{thm conjI}] @ @{thms rules}) 1)\<close>)
done
(*?x has 16, ?y has 32; rev(?y,?w) requires 561 (rather large) inferences
total inferences = 2 + 1 + 17 + 561 = 581*)
schematic_lemma "a:b:c:d:e:f:g:h:i:j:k:l:m:n:o:p:Nil = ?x & app(?x,?x,?y) & rev(?y,?w)"
-apply (tactic {*
- DEPTH_SOLVE (resolve_tac @{context} ([@{thm refl}, @{thm conjI}] @ @{thms rules}) 1) *})
+apply (tactic \<open>
+ DEPTH_SOLVE (resolve_tac @{context} ([@{thm refl}, @{thm conjI}] @ @{thms rules}) 1)\<close>)
done
end
--- a/src/FOL/ex/Propositional_Cla.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/FOL/ex/Propositional_Cla.thy Thu Jul 23 14:25:05 2015 +0200
@@ -3,13 +3,13 @@
Copyright 1991 University of Cambridge
*)
-section {* First-Order Logic: propositional examples (classical version) *}
+section \<open>First-Order Logic: propositional examples (classical version)\<close>
theory Propositional_Cla
imports FOL
begin
-text {* commutative laws of @{text "&"} and @{text "|"} *}
+text \<open>commutative laws of @{text "&"} and @{text "|"}\<close>
lemma "P & Q --> Q & P"
by (tactic "IntPr.fast_tac @{context} 1")
@@ -18,7 +18,7 @@
by fast
-text {* associative laws of @{text "&"} and @{text "|"} *}
+text \<open>associative laws of @{text "&"} and @{text "|"}\<close>
lemma "(P & Q) & R --> P & (Q & R)"
by fast
@@ -26,7 +26,7 @@
by fast
-text {* distributive laws of @{text "&"} and @{text "|"} *}
+text \<open>distributive laws of @{text "&"} and @{text "|"}\<close>
lemma "(P & Q) | R --> (P | R) & (Q | R)"
by fast
@@ -40,7 +40,7 @@
by fast
-text {* Laws involving implication *}
+text \<open>Laws involving implication\<close>
lemma "(P-->R) & (Q-->R) <-> (P|Q --> R)"
by fast
@@ -58,18 +58,18 @@
by fast
-text {* Propositions-as-types *}
+text \<open>Propositions-as-types\<close>
--- {* The combinator K *}
+-- \<open>The combinator K\<close>
lemma "P --> (Q --> P)"
by fast
--- {* The combinator S *}
+-- \<open>The combinator S\<close>
lemma "(P-->Q-->R) --> (P-->Q) --> (P-->R)"
by fast
--- {* Converse is classical *}
+-- \<open>Converse is classical\<close>
lemma "(P-->Q) | (P-->R) --> (P --> Q | R)"
by fast
@@ -77,7 +77,7 @@
by fast
-text {* Schwichtenberg's examples (via T. Nipkow) *}
+text \<open>Schwichtenberg's examples (via T. Nipkow)\<close>
lemma stab_imp: "(((Q-->R)-->R)-->Q) --> (((P-->Q)-->R)-->R)-->P-->Q"
by fast
--- a/src/FOL/ex/Propositional_Int.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/FOL/ex/Propositional_Int.thy Thu Jul 23 14:25:05 2015 +0200
@@ -3,13 +3,13 @@
Copyright 1991 University of Cambridge
*)
-section {* First-Order Logic: propositional examples (intuitionistic version) *}
+section \<open>First-Order Logic: propositional examples (intuitionistic version)\<close>
theory Propositional_Int
imports IFOL
begin
-text {* commutative laws of @{text "&"} and @{text "|"} *}
+text \<open>commutative laws of @{text "&"} and @{text "|"}\<close>
lemma "P & Q --> Q & P"
by (tactic "IntPr.fast_tac @{context} 1")
@@ -18,7 +18,7 @@
by (tactic "IntPr.fast_tac @{context} 1")
-text {* associative laws of @{text "&"} and @{text "|"} *}
+text \<open>associative laws of @{text "&"} and @{text "|"}\<close>
lemma "(P & Q) & R --> P & (Q & R)"
by (tactic "IntPr.fast_tac @{context} 1")
@@ -26,7 +26,7 @@
by (tactic "IntPr.fast_tac @{context} 1")
-text {* distributive laws of @{text "&"} and @{text "|"} *}
+text \<open>distributive laws of @{text "&"} and @{text "|"}\<close>
lemma "(P & Q) | R --> (P | R) & (Q | R)"
by (tactic "IntPr.fast_tac @{context} 1")
@@ -40,7 +40,7 @@
by (tactic "IntPr.fast_tac @{context} 1")
-text {* Laws involving implication *}
+text \<open>Laws involving implication\<close>
lemma "(P-->R) & (Q-->R) <-> (P|Q --> R)"
by (tactic "IntPr.fast_tac @{context} 1")
@@ -58,18 +58,18 @@
by (tactic "IntPr.fast_tac @{context} 1")
-text {* Propositions-as-types *}
+text \<open>Propositions-as-types\<close>
--- {* The combinator K *}
+-- \<open>The combinator K\<close>
lemma "P --> (Q --> P)"
by (tactic "IntPr.fast_tac @{context} 1")
--- {* The combinator S *}
+-- \<open>The combinator S\<close>
lemma "(P-->Q-->R) --> (P-->Q) --> (P-->R)"
by (tactic "IntPr.fast_tac @{context} 1")
--- {* Converse is classical *}
+-- \<open>Converse is classical\<close>
lemma "(P-->Q) | (P-->R) --> (P --> Q | R)"
by (tactic "IntPr.fast_tac @{context} 1")
@@ -77,7 +77,7 @@
by (tactic "IntPr.fast_tac @{context} 1")
-text {* Schwichtenberg's examples (via T. Nipkow) *}
+text \<open>Schwichtenberg's examples (via T. Nipkow)\<close>
lemma stab_imp: "(((Q-->R)-->R)-->Q) --> (((P-->Q)-->R)-->R)-->P-->Q"
by (tactic "IntPr.fast_tac @{context} 1")
--- a/src/FOL/ex/Quantifiers_Cla.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/FOL/ex/Quantifiers_Cla.thy Thu Jul 23 14:25:05 2015 +0200
@@ -3,7 +3,7 @@
Copyright 1991 University of Cambridge
*)
-section {* First-Order Logic: quantifier examples (classical version) *}
+section \<open>First-Order Logic: quantifier examples (classical version)\<close>
theory Quantifiers_Cla
imports FOL
@@ -16,7 +16,7 @@
by fast
--- {* Converse is false *}
+-- \<open>Converse is false\<close>
lemma "(ALL x. P(x)) | (ALL x. Q(x)) --> (ALL x. P(x) | Q(x))"
by fast
@@ -28,19 +28,19 @@
by fast
-text {* Some harder ones *}
+text \<open>Some harder ones\<close>
lemma "(EX x. P(x) | Q(x)) <-> (EX x. P(x)) | (EX x. Q(x))"
by fast
--- {* Converse is false *}
+-- \<open>Converse is false\<close>
lemma "(EX x. P(x)&Q(x)) --> (EX x. P(x)) & (EX x. Q(x))"
by fast
-text {* Basic test of quantifier reasoning *}
+text \<open>Basic test of quantifier reasoning\<close>
--- {* TRUE *}
+-- \<open>TRUE\<close>
lemma "(EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))"
by fast
@@ -48,7 +48,7 @@
by fast
-text {* The following should fail, as they are false! *}
+text \<open>The following should fail, as they are false!\<close>
lemma "(ALL x. EX y. Q(x,y)) --> (EX y. ALL x. Q(x,y))"
apply fast?
@@ -67,12 +67,12 @@
oops
-text {* Back to things that are provable \dots *}
+text \<open>Back to things that are provable \dots\<close>
lemma "(ALL x. P(x)-->Q(x)) & (EX x. P(x)) --> (EX x. Q(x))"
by fast
--- {* An example of why exI should be delayed as long as possible *}
+-- \<open>An example of why exI should be delayed as long as possible\<close>
lemma "(P --> (EX x. Q(x))) & P --> (EX x. Q(x))"
by fast
@@ -83,9 +83,9 @@
by fast
-text {* Some slow ones *}
+text \<open>Some slow ones\<close>
--- {* Principia Mathematica *11.53 *}
+-- \<open>Principia Mathematica *11.53\<close>
lemma "(ALL x y. P(x) --> Q(y)) <-> ((EX x. P(x)) --> (ALL y. Q(y)))"
by fast
--- a/src/FOL/ex/Quantifiers_Int.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/FOL/ex/Quantifiers_Int.thy Thu Jul 23 14:25:05 2015 +0200
@@ -3,7 +3,7 @@
Copyright 1991 University of Cambridge
*)
-section {* First-Order Logic: quantifier examples (intuitionistic version) *}
+section \<open>First-Order Logic: quantifier examples (intuitionistic version)\<close>
theory Quantifiers_Int
imports IFOL
@@ -16,7 +16,7 @@
by (tactic "IntPr.fast_tac @{context} 1")
--- {* Converse is false *}
+-- \<open>Converse is false\<close>
lemma "(ALL x. P(x)) | (ALL x. Q(x)) --> (ALL x. P(x) | Q(x))"
by (tactic "IntPr.fast_tac @{context} 1")
@@ -28,19 +28,19 @@
by (tactic "IntPr.fast_tac @{context} 1")
-text {* Some harder ones *}
+text \<open>Some harder ones\<close>
lemma "(EX x. P(x) | Q(x)) <-> (EX x. P(x)) | (EX x. Q(x))"
by (tactic "IntPr.fast_tac @{context} 1")
--- {* Converse is false *}
+-- \<open>Converse is false\<close>
lemma "(EX x. P(x)&Q(x)) --> (EX x. P(x)) & (EX x. Q(x))"
by (tactic "IntPr.fast_tac @{context} 1")
-text {* Basic test of quantifier reasoning *}
+text \<open>Basic test of quantifier reasoning\<close>
--- {* TRUE *}
+-- \<open>TRUE\<close>
lemma "(EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))"
by (tactic "IntPr.fast_tac @{context} 1")
@@ -48,7 +48,7 @@
by (tactic "IntPr.fast_tac @{context} 1")
-text {* The following should fail, as they are false! *}
+text \<open>The following should fail, as they are false!\<close>
lemma "(ALL x. EX y. Q(x,y)) --> (EX y. ALL x. Q(x,y))"
apply (tactic "IntPr.fast_tac @{context} 1")?
@@ -67,12 +67,12 @@
oops
-text {* Back to things that are provable \dots *}
+text \<open>Back to things that are provable \dots\<close>
lemma "(ALL x. P(x)-->Q(x)) & (EX x. P(x)) --> (EX x. Q(x))"
by (tactic "IntPr.fast_tac @{context} 1")
--- {* An example of why exI should be delayed as long as possible *}
+-- \<open>An example of why exI should be delayed as long as possible\<close>
lemma "(P --> (EX x. Q(x))) & P --> (EX x. Q(x))"
by (tactic "IntPr.fast_tac @{context} 1")
@@ -83,9 +83,9 @@
by (tactic "IntPr.fast_tac @{context} 1")
-text {* Some slow ones *}
+text \<open>Some slow ones\<close>
--- {* Principia Mathematica *11.53 *}
+-- \<open>Principia Mathematica *11.53\<close>
lemma "(ALL x y. P(x) --> Q(y)) <-> ((EX x. P(x)) --> (ALL y. Q(y)))"
by (tactic "IntPr.fast_tac @{context} 1")
--- a/src/FOLP/FOLP.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/FOLP/FOLP.thy Thu Jul 23 14:25:05 2015 +0200
@@ -3,7 +3,7 @@
Copyright 1992 University of Cambridge
*)
-section {* Classical First-Order Logic with Proofs *}
+section \<open>Classical First-Order Logic with Proofs\<close>
theory FOLP
imports IFOLP
@@ -56,8 +56,8 @@
and r2: "!!y. y:Q ==> g(y):R"
shows "?p : R"
apply (rule excluded_middle [THEN disjE])
- apply (tactic {* DEPTH_SOLVE (assume_tac @{context} 1 ORELSE
- resolve_tac @{context} [@{thm r1}, @{thm r2}, @{thm major} RS @{thm mp}] 1) *})
+ apply (tactic \<open>DEPTH_SOLVE (assume_tac @{context} 1 ORELSE
+ resolve_tac @{context} [@{thm r1}, @{thm r2}, @{thm major} RS @{thm mp}] 1)\<close>)
done
(*Double negation law*)
@@ -80,10 +80,10 @@
apply (insert major)
apply (unfold iff_def)
apply (rule conjE)
- apply (tactic {* DEPTH_SOLVE_1 (eresolve_tac @{context} @{thms impCE} 1 ORELSE
+ apply (tactic \<open>DEPTH_SOLVE_1 (eresolve_tac @{context} @{thms impCE} 1 ORELSE
eresolve_tac @{context} [@{thm notE}, @{thm impE}] 1 THEN assume_tac @{context} 1 ORELSE
assume_tac @{context} 1 ORELSE
- resolve_tac @{context} [@{thm r1}, @{thm r2}] 1) *})+
+ resolve_tac @{context} [@{thm r1}, @{thm r2}] 1)\<close>)+
done
@@ -101,7 +101,7 @@
ML_file "classical.ML" (* Patched because matching won't instantiate proof *)
ML_file "simp.ML" (* Patched because matching won't instantiate proof *)
-ML {*
+ML \<open>
structure Cla = Classical
(
val sizef = size_of_thm
@@ -128,14 +128,14 @@
val FOLP_dup_cs =
prop_cs addSIs [@{thm allI}] addIs [@{thm exCI}, @{thm ex1I}]
addSEs [@{thm exE}, @{thm ex1E}] addEs [@{thm all_dupE}];
-*}
+\<close>
schematic_lemma cla_rews:
"?p1 : P | ~P"
"?p2 : ~P | P"
"?p3 : ~ ~ P <-> P"
"?p4 : (~P --> P) <-> P"
- apply (tactic {* ALLGOALS (Cla.fast_tac @{context} FOLP_cs) *})
+ apply (tactic \<open>ALLGOALS (Cla.fast_tac @{context} FOLP_cs)\<close>)
done
ML_file "simpdata.ML"
--- a/src/FOLP/IFOLP.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/FOLP/IFOLP.thy Thu Jul 23 14:25:05 2015 +0200
@@ -3,7 +3,7 @@
Copyright 1992 University of Cambridge
*)
-section {* Intuitionistic First-Order Logic with Proofs *}
+section \<open>Intuitionistic First-Order Logic with Proofs\<close>
theory IFOLP
imports Pure
@@ -64,21 +64,21 @@
syntax "_Proof" :: "[p,o]=>prop" ("(_ /: _)" [51, 10] 5)
-parse_translation {*
+parse_translation \<open>
let fun proof_tr [p, P] = Const (@{const_syntax Proof}, dummyT) $ P $ p
in [(@{syntax_const "_Proof"}, K proof_tr)] end
-*}
+\<close>
(*show_proofs = true displays the proof terms -- they are ENORMOUS*)
-ML {* val show_proofs = Attrib.setup_config_bool @{binding show_proofs} (K false) *}
+ML \<open>val show_proofs = Attrib.setup_config_bool @{binding show_proofs} (K false)\<close>
-print_translation {*
+print_translation \<open>
let
fun proof_tr' ctxt [P, p] =
if Config.get ctxt show_proofs then Const (@{syntax_const "_Proof"}, dummyT) $ p $ P
else P
in [(@{const_syntax Proof}, proof_tr')] end
-*}
+\<close>
(**** Propositional logic ****)
@@ -248,7 +248,7 @@
Fails unless one assumption is equal and exactly one is unifiable
**)
-ML {*
+ML \<open>
local
fun discard_proof (Const (@{const_name Proof}, _) $ P $ _) = P;
in
@@ -265,23 +265,23 @@
else no_tac
end);
end;
-*}
+\<close>
(*** Modus Ponens Tactics ***)
(*Finds P-->Q and P in the assumptions, replaces implication by Q *)
-ML {*
+ML \<open>
fun mp_tac ctxt i =
eresolve_tac ctxt [@{thm notE}, make_elim @{thm mp}] i THEN assume_tac ctxt i
-*}
+\<close>
method_setup mp = \<open>Scan.succeed (SIMPLE_METHOD' o mp_tac)\<close>
(*Like mp_tac but instantiates no variables*)
-ML {*
+ML \<open>
fun int_uniq_mp_tac ctxt i =
eresolve_tac ctxt [@{thm notE}, @{thm impE}] i THEN uniq_assume_tac ctxt i
-*}
+\<close>
(*** If-and-only-if ***)
@@ -360,11 +360,11 @@
(*** <-> congruence rules for simplification ***)
(*Use iffE on a premise. For conj_cong, imp_cong, all_cong, ex_cong*)
-ML {*
+ML \<open>
fun iff_tac ctxt prems i =
resolve_tac ctxt (prems RL [@{thm iffE}]) i THEN
REPEAT1 (eresolve_tac ctxt [asm_rl, @{thm mp}] i)
-*}
+\<close>
method_setup iff =
\<open>Attrib.thms >> (fn prems => fn ctxt => SIMPLE_METHOD' (iff_tac ctxt prems))\<close>
@@ -503,20 +503,20 @@
schematic_lemma pred1_cong: "p:a=a' ==> ?p:P(a) <-> P(a')"
apply (rule iffI)
- apply (tactic {*
- DEPTH_SOLVE (assume_tac @{context} 1 ORELSE eresolve_tac @{context} [@{thm subst}, @{thm ssubst}] 1) *})
+ apply (tactic \<open>
+ DEPTH_SOLVE (assume_tac @{context} 1 ORELSE eresolve_tac @{context} [@{thm subst}, @{thm ssubst}] 1)\<close>)
done
schematic_lemma pred2_cong: "[| p:a=a'; q:b=b' |] ==> ?p:P(a,b) <-> P(a',b')"
apply (rule iffI)
- apply (tactic {*
- DEPTH_SOLVE (assume_tac @{context} 1 ORELSE eresolve_tac @{context} [@{thm subst}, @{thm ssubst}] 1) *})
+ apply (tactic \<open>
+ DEPTH_SOLVE (assume_tac @{context} 1 ORELSE eresolve_tac @{context} [@{thm subst}, @{thm ssubst}] 1)\<close>)
done
schematic_lemma pred3_cong: "[| p:a=a'; q:b=b'; r:c=c' |] ==> ?p:P(a,b,c) <-> P(a',b',c')"
apply (rule iffI)
- apply (tactic {*
- DEPTH_SOLVE (assume_tac @{context} 1 ORELSE eresolve_tac @{context} [@{thm subst}, @{thm ssubst}] 1) *})
+ apply (tactic \<open>
+ DEPTH_SOLVE (assume_tac @{context} 1 ORELSE eresolve_tac @{context} [@{thm subst}, @{thm ssubst}] 1)\<close>)
done
lemmas pred_congs = pred1_cong pred2_cong pred3_cong
@@ -543,9 +543,9 @@
assumes major: "p:(P|Q)-->S"
and minor: "!!x y.[| x:P-->S; y:Q-->S |] ==> q(x,y):R"
shows "?p:R"
- apply (tactic {* DEPTH_SOLVE (assume_tac @{context} 1 ORELSE
+ apply (tactic \<open>DEPTH_SOLVE (assume_tac @{context} 1 ORELSE
resolve_tac @{context} [@{thm disjI1}, @{thm disjI2}, @{thm impI},
- @{thm major} RS @{thm mp}, @{thm minor}] 1) *})
+ @{thm major} RS @{thm mp}, @{thm minor}] 1)\<close>)
done
(*Simplifies the implication. Classical version is stronger.
@@ -607,7 +607,7 @@
ML_file "hypsubst.ML"
-ML {*
+ML \<open>
structure Hypsubst = Hypsubst
(
(*Take apart an equality judgement; otherwise raise Match!*)
@@ -625,7 +625,7 @@
val thin_refl = @{thm thin_refl}
);
open Hypsubst;
-*}
+\<close>
ML_file "intprover.ML"
@@ -641,7 +641,7 @@
"?p6 : P & ~P <-> False"
"?p7 : ~P & P <-> False"
"?p8 : (P & Q) & R <-> P & (Q & R)"
- apply (tactic {* fn st => IntPr.fast_tac @{context} 1 st *})+
+ apply (tactic \<open>fn st => IntPr.fast_tac @{context} 1 st\<close>)+
done
schematic_lemma disj_rews:
@@ -651,13 +651,13 @@
"?p4 : False | P <-> P"
"?p5 : P | P <-> P"
"?p6 : (P | Q) | R <-> P | (Q | R)"
- apply (tactic {* IntPr.fast_tac @{context} 1 *})+
+ apply (tactic \<open>IntPr.fast_tac @{context} 1\<close>)+
done
schematic_lemma not_rews:
"?p1 : ~ False <-> True"
"?p2 : ~ True <-> False"
- apply (tactic {* IntPr.fast_tac @{context} 1 *})+
+ apply (tactic \<open>IntPr.fast_tac @{context} 1\<close>)+
done
schematic_lemma imp_rews:
@@ -667,7 +667,7 @@
"?p4 : (True --> P) <-> P"
"?p5 : (P --> P) <-> True"
"?p6 : (P --> ~P) <-> ~P"
- apply (tactic {* IntPr.fast_tac @{context} 1 *})+
+ apply (tactic \<open>IntPr.fast_tac @{context} 1\<close>)+
done
schematic_lemma iff_rews:
@@ -676,13 +676,13 @@
"?p3 : (P <-> P) <-> True"
"?p4 : (False <-> P) <-> ~P"
"?p5 : (P <-> False) <-> ~P"
- apply (tactic {* IntPr.fast_tac @{context} 1 *})+
+ apply (tactic \<open>IntPr.fast_tac @{context} 1\<close>)+
done
schematic_lemma quant_rews:
"?p1 : (ALL x. P) <-> P"
"?p2 : (EX x. P) <-> P"
- apply (tactic {* IntPr.fast_tac @{context} 1 *})+
+ apply (tactic \<open>IntPr.fast_tac @{context} 1\<close>)+
done
(*These are NOT supplied by default!*)
@@ -691,7 +691,7 @@
"?p2 : P & (Q | R) <-> P&Q | P&R"
"?p3 : (Q | R) & P <-> Q&P | R&P"
"?p4 : (P | Q --> R) <-> (P --> R) & (Q --> R)"
- apply (tactic {* IntPr.fast_tac @{context} 1 *})+
+ apply (tactic \<open>IntPr.fast_tac @{context} 1\<close>)+
done
schematic_lemma distrib_rews2:
@@ -699,17 +699,17 @@
"?p2 : ((EX x. NORM(P(x))) --> Q) <-> (ALL x. NORM(P(x)) --> Q)"
"?p3 : (EX x. NORM(P(x))) & NORM(Q) <-> (EX x. NORM(P(x)) & NORM(Q))"
"?p4 : NORM(Q) & (EX x. NORM(P(x))) <-> (EX x. NORM(Q) & NORM(P(x)))"
- apply (tactic {* IntPr.fast_tac @{context} 1 *})+
+ apply (tactic \<open>IntPr.fast_tac @{context} 1\<close>)+
done
lemmas distrib_rews = distrib_rews1 distrib_rews2
schematic_lemma P_Imp_P_iff_T: "p:P ==> ?p:(P <-> True)"
- apply (tactic {* IntPr.fast_tac @{context} 1 *})
+ apply (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
done
schematic_lemma not_P_imp_P_iff_F: "p:~P ==> ?p:(P <-> False)"
- apply (tactic {* IntPr.fast_tac @{context} 1 *})
+ apply (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
done
end
--- a/src/FOLP/ex/Classical.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/FOLP/ex/Classical.thy Thu Jul 23 14:25:05 2015 +0200
@@ -288,7 +288,7 @@
text "Problem 58 NOT PROVED AUTOMATICALLY"
schematic_lemma "?p : (ALL x y. f(x)=g(y)) --> (ALL x y. f(f(x))=f(g(y)))"
supply f_cong = subst_context [where t = f]
- by (tactic {* fast_tac @{context} (FOLP_cs addSIs [@{thm f_cong}]) 1 *})
+ by (tactic \<open>fast_tac @{context} (FOLP_cs addSIs [@{thm f_cong}]) 1\<close>)
text "Problem 59"
schematic_lemma "?p : (ALL x. P(x) <-> ~P(f(x))) --> (EX x. P(x) & ~P(f(x)))"
--- a/src/FOLP/ex/Foundation.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/FOLP/ex/Foundation.thy Thu Jul 23 14:25:05 2015 +0200
@@ -18,7 +18,7 @@
apply assumption
done
-text {*A form of conj-elimination*}
+text \<open>A form of conj-elimination\<close>
schematic_lemma
assumes "p : A & B"
and "!!x y. x : A ==> y : B ==> f(x, y) : C"
@@ -99,7 +99,7 @@
apply (rule refl)?
oops
-text {* Parallel lifting example. *}
+text \<open>Parallel lifting example.\<close>
schematic_lemma "?p : EX u. ALL x. EX v. ALL y. EX w. P(u,x,v,y,w)"
apply (rule exI allI)
apply (rule exI allI)
@@ -121,7 +121,7 @@
apply assumption
done
-text {* A bigger demonstration of quantifiers -- not in the paper. *}
+text \<open>A bigger demonstration of quantifiers -- not in the paper.\<close>
schematic_lemma "?p : (EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))"
apply (rule impI)
apply (rule allI)
--- a/src/FOLP/ex/If.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/FOLP/ex/If.thy Thu Jul 23 14:25:05 2015 +0200
@@ -9,7 +9,7 @@
assumes "!!x. x : P ==> f(x) : Q" "!!x. x : ~P ==> g(x) : R"
shows "?p : if(P,Q,R)"
apply (unfold if_def)
-apply (tactic {* fast_tac @{context} (FOLP_cs addIs @{thms assms}) 1 *})
+apply (tactic \<open>fast_tac @{context} (FOLP_cs addIs @{thms assms}) 1\<close>)
done
schematic_lemma ifE:
@@ -19,7 +19,7 @@
shows "?p : S"
apply (insert 1)
apply (unfold if_def)
-apply (tactic {* fast_tac @{context} (FOLP_cs addIs [@{thm 2}, @{thm 3}]) 1 *})
+apply (tactic \<open>fast_tac @{context} (FOLP_cs addIs [@{thm 2}, @{thm 3}]) 1\<close>)
done
schematic_lemma if_commute: "?p : if(P, if(Q,A,B), if(Q,C,D)) <-> if(Q, if(P,A,C), if(P,B,D))"
@@ -30,14 +30,14 @@
apply (rule ifI)
oops
-ML {* val if_cs = FOLP_cs addSIs [@{thm ifI}] addSEs [@{thm ifE}] *}
+ML \<open>val if_cs = FOLP_cs addSIs [@{thm ifI}] addSEs [@{thm ifE}]\<close>
schematic_lemma if_commute: "?p : if(P, if(Q,A,B), if(Q,C,D)) <-> if(Q, if(P,A,C), if(P,B,D))"
-apply (tactic {* fast_tac @{context} if_cs 1 *})
+apply (tactic \<open>fast_tac @{context} if_cs 1\<close>)
done
schematic_lemma nested_ifs: "?p : if(if(P,Q,R), A, B) <-> if(P, if(Q,A,B), if(R,A,B))"
-apply (tactic {* fast_tac @{context} if_cs 1 *})
+apply (tactic \<open>fast_tac @{context} if_cs 1\<close>)
done
end
--- a/src/FOLP/ex/Intro.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/FOLP/ex/Intro.thy Thu Jul 23 14:25:05 2015 +0200
@@ -5,13 +5,13 @@
Derives some inference rules, illustrating the use of definitions.
*)
-section {* Examples for the manual ``Introduction to Isabelle'' *}
+section \<open>Examples for the manual ``Introduction to Isabelle''\<close>
theory Intro
imports FOLP
begin
-subsubsection {* Some simple backward proofs *}
+subsubsection \<open>Some simple backward proofs\<close>
schematic_lemma mythm: "?p : P|P --> P"
apply (rule impI)
@@ -41,21 +41,21 @@
done
-subsubsection {* Demonstration of @{text "fast"} *}
+subsubsection \<open>Demonstration of @{text "fast"}\<close>
schematic_lemma "?p : (EX y. ALL x. J(y,x) <-> ~J(x,x))
--> ~ (ALL x. EX y. ALL z. J(z,y) <-> ~ J(z,x))"
-apply (tactic {* fast_tac @{context} FOLP_cs 1 *})
+apply (tactic \<open>fast_tac @{context} FOLP_cs 1\<close>)
done
schematic_lemma "?p : ALL x. P(x,f(x)) <->
(EX y. (ALL z. P(z,y) --> P(z,f(x))) & P(x,y))"
-apply (tactic {* fast_tac @{context} FOLP_cs 1 *})
+apply (tactic \<open>fast_tac @{context} FOLP_cs 1\<close>)
done
-subsubsection {* Derivation of conjunction elimination rule *}
+subsubsection \<open>Derivation of conjunction elimination rule\<close>
schematic_lemma
assumes major: "p : P&Q"
@@ -67,9 +67,9 @@
done
-subsection {* Derived rules involving definitions *}
+subsection \<open>Derived rules involving definitions\<close>
-text {* Derivation of negation introduction *}
+text \<open>Derivation of negation introduction\<close>
schematic_lemma
assumes "!!x. x : P ==> f(x) : False"
@@ -90,7 +90,7 @@
apply (rule minor)
done
-text {* Alternative proof of the result above *}
+text \<open>Alternative proof of the result above\<close>
schematic_lemma
assumes major: "p : ~P"
and minor: "q : P"
--- a/src/FOLP/ex/Intuitionistic.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/FOLP/ex/Intuitionistic.thy Thu Jul 23 14:25:05 2015 +0200
@@ -31,167 +31,167 @@
begin
schematic_lemma "?p : ~~(P&Q) <-> ~~P & ~~Q"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
schematic_lemma "?p : ~~~P <-> ~P"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
schematic_lemma "?p : ~~((P --> Q | R) --> (P-->Q) | (P-->R))"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
schematic_lemma "?p : (P<->Q) <-> (Q<->P)"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
-subsection {* Lemmas for the propositional double-negation translation *}
+subsection \<open>Lemmas for the propositional double-negation translation\<close>
schematic_lemma "?p : P --> ~~P"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
schematic_lemma "?p : ~~(~~P --> P)"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
schematic_lemma "?p : ~~P & ~~(P --> Q) --> ~~Q"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
-subsection {* The following are classically but not constructively valid *}
+subsection \<open>The following are classically but not constructively valid\<close>
(*The attempt to prove them terminates quickly!*)
schematic_lemma "?p : ((P-->Q) --> P) --> P"
- apply (tactic {* IntPr.fast_tac @{context} 1 *})?
+ apply (tactic \<open>IntPr.fast_tac @{context} 1\<close>)?
oops
schematic_lemma "?p : (P&Q-->R) --> (P-->R) | (Q-->R)"
- apply (tactic {* IntPr.fast_tac @{context} 1 *})?
+ apply (tactic \<open>IntPr.fast_tac @{context} 1\<close>)?
oops
-subsection {* Intuitionistic FOL: propositional problems based on Pelletier *}
+subsection \<open>Intuitionistic FOL: propositional problems based on Pelletier\<close>
text "Problem ~~1"
schematic_lemma "?p : ~~((P-->Q) <-> (~Q --> ~P))"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text "Problem ~~2"
schematic_lemma "?p : ~~(~~P <-> P)"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text "Problem 3"
schematic_lemma "?p : ~(P-->Q) --> (Q-->P)"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text "Problem ~~4"
schematic_lemma "?p : ~~((~P-->Q) <-> (~Q --> P))"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text "Problem ~~5"
schematic_lemma "?p : ~~((P|Q-->P|R) --> P|(Q-->R))"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text "Problem ~~6"
schematic_lemma "?p : ~~(P | ~P)"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text "Problem ~~7"
schematic_lemma "?p : ~~(P | ~~~P)"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text "Problem ~~8. Peirce's law"
schematic_lemma "?p : ~~(((P-->Q) --> P) --> P)"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text "Problem 9"
schematic_lemma "?p : ((P|Q) & (~P|Q) & (P| ~Q)) --> ~ (~P | ~Q)"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text "Problem 10"
schematic_lemma "?p : (Q-->R) --> (R-->P&Q) --> (P-->(Q|R)) --> (P<->Q)"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text "11. Proved in each direction (incorrectly, says Pelletier!!) "
schematic_lemma "?p : P<->P"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text "Problem ~~12. Dijkstra's law "
schematic_lemma "?p : ~~(((P <-> Q) <-> R) <-> (P <-> (Q <-> R)))"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
schematic_lemma "?p : ((P <-> Q) <-> R) --> ~~(P <-> (Q <-> R))"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text "Problem 13. Distributive law"
schematic_lemma "?p : P | (Q & R) <-> (P | Q) & (P | R)"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text "Problem ~~14"
schematic_lemma "?p : ~~((P <-> Q) <-> ((Q | ~P) & (~Q|P)))"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text "Problem ~~15"
schematic_lemma "?p : ~~((P --> Q) <-> (~P | Q))"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text "Problem ~~16"
schematic_lemma "?p : ~~((P-->Q) | (Q-->P))"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text "Problem ~~17"
schematic_lemma "?p : ~~(((P & (Q-->R))-->S) <-> ((~P | Q | S) & (~P | ~R | S)))"
- by (tactic {* IntPr.fast_tac @{context} 1 *}) -- slow
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>) -- slow
-subsection {* Examples with quantifiers *}
+subsection \<open>Examples with quantifiers\<close>
text "The converse is classical in the following implications..."
schematic_lemma "?p : (EX x. P(x)-->Q) --> (ALL x. P(x)) --> Q"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
schematic_lemma "?p : ((ALL x. P(x))-->Q) --> ~ (ALL x. P(x) & ~Q)"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
schematic_lemma "?p : ((ALL x. ~P(x))-->Q) --> ~ (ALL x. ~ (P(x)|Q))"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
schematic_lemma "?p : (ALL x. P(x)) | Q --> (ALL x. P(x) | Q)"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
schematic_lemma "?p : (EX x. P --> Q(x)) --> (P --> (EX x. Q(x)))"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text "The following are not constructively valid!"
text "The attempt to prove them terminates quickly!"
schematic_lemma "?p : ((ALL x. P(x))-->Q) --> (EX x. P(x)-->Q)"
- apply (tactic {* IntPr.fast_tac @{context} 1 *})?
+ apply (tactic \<open>IntPr.fast_tac @{context} 1\<close>)?
oops
schematic_lemma "?p : (P --> (EX x. Q(x))) --> (EX x. P-->Q(x))"
- apply (tactic {* IntPr.fast_tac @{context} 1 *})?
+ apply (tactic \<open>IntPr.fast_tac @{context} 1\<close>)?
oops
schematic_lemma "?p : (ALL x. P(x) | Q) --> ((ALL x. P(x)) | Q)"
- apply (tactic {* IntPr.fast_tac @{context} 1 *})?
+ apply (tactic \<open>IntPr.fast_tac @{context} 1\<close>)?
oops
schematic_lemma "?p : (ALL x. ~~P(x)) --> ~~(ALL x. P(x))"
- apply (tactic {* IntPr.fast_tac @{context} 1 *})?
+ apply (tactic \<open>IntPr.fast_tac @{context} 1\<close>)?
oops
(*Classically but not intuitionistically valid. Proved by a bug in 1986!*)
schematic_lemma "?p : EX x. Q(x) --> (ALL x. Q(x))"
- apply (tactic {* IntPr.fast_tac @{context} 1 *})?
+ apply (tactic \<open>IntPr.fast_tac @{context} 1\<close>)?
oops
subsection "Hard examples with quantifiers"
-text {*
+text \<open>
The ones that have not been proved are not known to be valid!
Some will require quantifier duplication -- not currently available.
-*}
+\<close>
text "Problem ~~18"
schematic_lemma "?p : ~~(EX y. ALL x. P(y)-->P(x))" oops
@@ -204,7 +204,7 @@
text "Problem 20"
schematic_lemma "?p : (ALL x y. EX z. ALL w. (P(x)&Q(y)-->R(z)&S(w)))
--> (EX x y. P(x) & Q(y)) --> (EX z. R(z))"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text "Problem 21"
schematic_lemma "?p : (EX x. P-->Q(x)) & (EX x. Q(x)-->P) --> ~~(EX x. P<->Q(x))" oops
@@ -212,11 +212,11 @@
text "Problem 22"
schematic_lemma "?p : (ALL x. P <-> Q(x)) --> (P <-> (ALL x. Q(x)))"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text "Problem ~~23"
schematic_lemma "?p : ~~ ((ALL x. P | Q(x)) <-> (P | (ALL x. Q(x))))"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text "Problem 24"
schematic_lemma "?p : ~(EX x. S(x)&Q(x)) & (ALL x. P(x) --> Q(x)|R(x)) &
@@ -287,7 +287,7 @@
schematic_lemma
"?p : (EX z w. ALL x y. P(x,y) <-> (x=z & y=w)) -->
(EX z. ALL x. EX w. (ALL y. P(x,y) <-> y=w) <-> x=z)"
- by (tactic "IntPr.best_tac @{context} 1") -- {*60 seconds*}
+ by (tactic "IntPr.best_tac @{context} 1") -- \<open>60 seconds\<close>
text "Problem 56"
schematic_lemma "?p : (ALL x. (EX y. P(y) & x=f(y)) --> P(x)) <-> (ALL x. P(x) --> P(f(x)))"
--- a/src/FOLP/ex/Nat.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/FOLP/ex/Nat.thy Thu Jul 23 14:25:05 2015 +0200
@@ -3,7 +3,7 @@
Copyright 1992 University of Cambridge
*)
-section {* Theory of the natural numbers: Peano's axioms, primitive recursion *}
+section \<open>Theory of the natural numbers: Peano's axioms, primitive recursion\<close>
theory Nat
imports FOLP
@@ -38,7 +38,7 @@
where "m + n == rec(m, n, %x y. Suc(y))"
-subsection {* Proofs about the natural numbers *}
+subsection \<open>Proofs about the natural numbers\<close>
schematic_lemma Suc_n_not_n: "?p : ~ (Suc(k) = k)"
apply (rule_tac n = k in induct)
@@ -81,27 +81,27 @@
lemmas nat_congs = Suc_cong Plus_cong
-ML {*
+ML \<open>
val add_ss =
FOLP_ss addcongs @{thms nat_congs}
|> fold (addrew @{context}) @{thms add_0 add_Suc}
-*}
+\<close>
schematic_lemma add_assoc: "?p : (k+m)+n = k+(m+n)"
apply (rule_tac n = k in induct)
-apply (tactic {* SIMP_TAC @{context} add_ss 1 *})
-apply (tactic {* ASM_SIMP_TAC @{context} add_ss 1 *})
+apply (tactic \<open>SIMP_TAC @{context} add_ss 1\<close>)
+apply (tactic \<open>ASM_SIMP_TAC @{context} add_ss 1\<close>)
done
schematic_lemma add_0_right: "?p : m+0 = m"
apply (rule_tac n = m in induct)
-apply (tactic {* SIMP_TAC @{context} add_ss 1 *})
-apply (tactic {* ASM_SIMP_TAC @{context} add_ss 1 *})
+apply (tactic \<open>SIMP_TAC @{context} add_ss 1\<close>)
+apply (tactic \<open>ASM_SIMP_TAC @{context} add_ss 1\<close>)
done
schematic_lemma add_Suc_right: "?p : m+Suc(n) = Suc(m+n)"
apply (rule_tac n = m in induct)
-apply (tactic {* ALLGOALS (ASM_SIMP_TAC @{context} add_ss) *})
+apply (tactic \<open>ALLGOALS (ASM_SIMP_TAC @{context} add_ss)\<close>)
done
(*mk_typed_congs appears not to work with FOLP's version of subst*)
--- a/src/FOLP/ex/Propositional_Cla.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/FOLP/ex/Propositional_Cla.thy Thu Jul 23 14:25:05 2015 +0200
@@ -3,7 +3,7 @@
Copyright 1991 University of Cambridge
*)
-section {* First-Order Logic: propositional examples *}
+section \<open>First-Order Logic: propositional examples\<close>
theory Propositional_Cla
imports FOLP
@@ -12,106 +12,106 @@
text "commutative laws of & and | "
schematic_lemma "?p : P & Q --> Q & P"
- by (tactic {* Cla.fast_tac @{context} FOLP_cs 1 *})
+ by (tactic \<open>Cla.fast_tac @{context} FOLP_cs 1\<close>)
schematic_lemma "?p : P | Q --> Q | P"
- by (tactic {* Cla.fast_tac @{context} FOLP_cs 1 *})
+ by (tactic \<open>Cla.fast_tac @{context} FOLP_cs 1\<close>)
text "associative laws of & and | "
schematic_lemma "?p : (P & Q) & R --> P & (Q & R)"
- by (tactic {* Cla.fast_tac @{context} FOLP_cs 1 *})
+ by (tactic \<open>Cla.fast_tac @{context} FOLP_cs 1\<close>)
schematic_lemma "?p : (P | Q) | R --> P | (Q | R)"
- by (tactic {* Cla.fast_tac @{context} FOLP_cs 1 *})
+ by (tactic \<open>Cla.fast_tac @{context} FOLP_cs 1\<close>)
text "distributive laws of & and | "
schematic_lemma "?p : (P & Q) | R --> (P | R) & (Q | R)"
- by (tactic {* Cla.fast_tac @{context} FOLP_cs 1 *})
+ by (tactic \<open>Cla.fast_tac @{context} FOLP_cs 1\<close>)
schematic_lemma "?p : (P | R) & (Q | R) --> (P & Q) | R"
- by (tactic {* Cla.fast_tac @{context} FOLP_cs 1 *})
+ by (tactic \<open>Cla.fast_tac @{context} FOLP_cs 1\<close>)
schematic_lemma "?p : (P | Q) & R --> (P & R) | (Q & R)"
- by (tactic {* Cla.fast_tac @{context} FOLP_cs 1 *})
+ by (tactic \<open>Cla.fast_tac @{context} FOLP_cs 1\<close>)
schematic_lemma "?p : (P & R) | (Q & R) --> (P | Q) & R"
- by (tactic {* Cla.fast_tac @{context} FOLP_cs 1 *})
+ by (tactic \<open>Cla.fast_tac @{context} FOLP_cs 1\<close>)
text "Laws involving implication"
schematic_lemma "?p : (P-->R) & (Q-->R) <-> (P|Q --> R)"
- by (tactic {* Cla.fast_tac @{context} FOLP_cs 1 *})
+ by (tactic \<open>Cla.fast_tac @{context} FOLP_cs 1\<close>)
schematic_lemma "?p : (P & Q --> R) <-> (P--> (Q-->R))"
- by (tactic {* Cla.fast_tac @{context} FOLP_cs 1 *})
+ by (tactic \<open>Cla.fast_tac @{context} FOLP_cs 1\<close>)
schematic_lemma "?p : ((P-->R)-->R) --> ((Q-->R)-->R) --> (P&Q-->R) --> R"
- by (tactic {* Cla.fast_tac @{context} FOLP_cs 1 *})
+ by (tactic \<open>Cla.fast_tac @{context} FOLP_cs 1\<close>)
schematic_lemma "?p : ~(P-->R) --> ~(Q-->R) --> ~(P&Q-->R)"
- by (tactic {* Cla.fast_tac @{context} FOLP_cs 1 *})
+ by (tactic \<open>Cla.fast_tac @{context} FOLP_cs 1\<close>)
schematic_lemma "?p : (P --> Q & R) <-> (P-->Q) & (P-->R)"
- by (tactic {* Cla.fast_tac @{context} FOLP_cs 1 *})
+ by (tactic \<open>Cla.fast_tac @{context} FOLP_cs 1\<close>)
text "Propositions-as-types"
(*The combinator K*)
schematic_lemma "?p : P --> (Q --> P)"
- by (tactic {* Cla.fast_tac @{context} FOLP_cs 1 *})
+ by (tactic \<open>Cla.fast_tac @{context} FOLP_cs 1\<close>)
(*The combinator S*)
schematic_lemma "?p : (P-->Q-->R) --> (P-->Q) --> (P-->R)"
- by (tactic {* Cla.fast_tac @{context} FOLP_cs 1 *})
+ by (tactic \<open>Cla.fast_tac @{context} FOLP_cs 1\<close>)
(*Converse is classical*)
schematic_lemma "?p : (P-->Q) | (P-->R) --> (P --> Q | R)"
- by (tactic {* Cla.fast_tac @{context} FOLP_cs 1 *})
+ by (tactic \<open>Cla.fast_tac @{context} FOLP_cs 1\<close>)
schematic_lemma "?p : (P-->Q) --> (~Q --> ~P)"
- by (tactic {* Cla.fast_tac @{context} FOLP_cs 1 *})
+ by (tactic \<open>Cla.fast_tac @{context} FOLP_cs 1\<close>)
text "Schwichtenberg's examples (via T. Nipkow)"
schematic_lemma stab_imp: "?p : (((Q-->R)-->R)-->Q) --> (((P-->Q)-->R)-->R)-->P-->Q"
- by (tactic {* Cla.fast_tac @{context} FOLP_cs 1 *})
+ by (tactic \<open>Cla.fast_tac @{context} FOLP_cs 1\<close>)
schematic_lemma stab_to_peirce: "?p : (((P --> R) --> R) --> P) --> (((Q --> R) --> R) --> Q)
--> ((P --> Q) --> P) --> P"
- by (tactic {* Cla.fast_tac @{context} FOLP_cs 1 *})
+ by (tactic \<open>Cla.fast_tac @{context} FOLP_cs 1\<close>)
schematic_lemma peirce_imp1: "?p : (((Q --> R) --> Q) --> Q)
--> (((P --> Q) --> R) --> P --> Q) --> P --> Q"
- by (tactic {* Cla.fast_tac @{context} FOLP_cs 1 *})
+ by (tactic \<open>Cla.fast_tac @{context} FOLP_cs 1\<close>)
schematic_lemma peirce_imp2: "?p : (((P --> R) --> P) --> P) --> ((P --> Q --> R) --> P) --> P"
- by (tactic {* Cla.fast_tac @{context} FOLP_cs 1 *})
+ by (tactic \<open>Cla.fast_tac @{context} FOLP_cs 1\<close>)
schematic_lemma mints: "?p : ((((P --> Q) --> P) --> P) --> Q) --> Q"
- by (tactic {* Cla.fast_tac @{context} FOLP_cs 1 *})
+ by (tactic \<open>Cla.fast_tac @{context} FOLP_cs 1\<close>)
schematic_lemma mints_solovev: "?p : (P --> (Q --> R) --> Q) --> ((P --> Q) --> R) --> R"
- by (tactic {* Cla.fast_tac @{context} FOLP_cs 1 *})
+ by (tactic \<open>Cla.fast_tac @{context} FOLP_cs 1\<close>)
schematic_lemma tatsuta: "?p : (((P7 --> P1) --> P10) --> P4 --> P5)
--> (((P8 --> P2) --> P9) --> P3 --> P10)
--> (P1 --> P8) --> P6 --> P7
--> (((P3 --> P2) --> P9) --> P4)
--> (P1 --> P3) --> (((P6 --> P1) --> P2) --> P9) --> P5"
- by (tactic {* Cla.fast_tac @{context} FOLP_cs 1 *})
+ by (tactic \<open>Cla.fast_tac @{context} FOLP_cs 1\<close>)
schematic_lemma tatsuta1: "?p : (((P8 --> P2) --> P9) --> P3 --> P10)
--> (((P3 --> P2) --> P9) --> P4)
--> (((P6 --> P1) --> P2) --> P9)
--> (((P7 --> P1) --> P10) --> P4 --> P5)
--> (P1 --> P3) --> (P1 --> P8) --> P6 --> P7 --> P5"
- by (tactic {* Cla.fast_tac @{context} FOLP_cs 1 *})
+ by (tactic \<open>Cla.fast_tac @{context} FOLP_cs 1\<close>)
end
--- a/src/FOLP/ex/Propositional_Int.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/FOLP/ex/Propositional_Int.thy Thu Jul 23 14:25:05 2015 +0200
@@ -3,7 +3,7 @@
Copyright 1991 University of Cambridge
*)
-section {* First-Order Logic: propositional examples *}
+section \<open>First-Order Logic: propositional examples\<close>
theory Propositional_Int
imports IFOLP
@@ -12,106 +12,106 @@
text "commutative laws of & and | "
schematic_lemma "?p : P & Q --> Q & P"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
schematic_lemma "?p : P | Q --> Q | P"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text "associative laws of & and | "
schematic_lemma "?p : (P & Q) & R --> P & (Q & R)"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
schematic_lemma "?p : (P | Q) | R --> P | (Q | R)"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text "distributive laws of & and | "
schematic_lemma "?p : (P & Q) | R --> (P | R) & (Q | R)"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
schematic_lemma "?p : (P | R) & (Q | R) --> (P & Q) | R"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
schematic_lemma "?p : (P | Q) & R --> (P & R) | (Q & R)"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
schematic_lemma "?p : (P & R) | (Q & R) --> (P | Q) & R"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text "Laws involving implication"
schematic_lemma "?p : (P-->R) & (Q-->R) <-> (P|Q --> R)"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
schematic_lemma "?p : (P & Q --> R) <-> (P--> (Q-->R))"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
schematic_lemma "?p : ((P-->R)-->R) --> ((Q-->R)-->R) --> (P&Q-->R) --> R"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
schematic_lemma "?p : ~(P-->R) --> ~(Q-->R) --> ~(P&Q-->R)"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
schematic_lemma "?p : (P --> Q & R) <-> (P-->Q) & (P-->R)"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text "Propositions-as-types"
(*The combinator K*)
schematic_lemma "?p : P --> (Q --> P)"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
(*The combinator S*)
schematic_lemma "?p : (P-->Q-->R) --> (P-->Q) --> (P-->R)"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
(*Converse is classical*)
schematic_lemma "?p : (P-->Q) | (P-->R) --> (P --> Q | R)"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
schematic_lemma "?p : (P-->Q) --> (~Q --> ~P)"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text "Schwichtenberg's examples (via T. Nipkow)"
schematic_lemma stab_imp: "?p : (((Q-->R)-->R)-->Q) --> (((P-->Q)-->R)-->R)-->P-->Q"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
schematic_lemma stab_to_peirce: "?p : (((P --> R) --> R) --> P) --> (((Q --> R) --> R) --> Q)
--> ((P --> Q) --> P) --> P"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
schematic_lemma peirce_imp1: "?p : (((Q --> R) --> Q) --> Q)
--> (((P --> Q) --> R) --> P --> Q) --> P --> Q"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
schematic_lemma peirce_imp2: "?p : (((P --> R) --> P) --> P) --> ((P --> Q --> R) --> P) --> P"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
schematic_lemma mints: "?p : ((((P --> Q) --> P) --> P) --> Q) --> Q"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
schematic_lemma mints_solovev: "?p : (P --> (Q --> R) --> Q) --> ((P --> Q) --> R) --> R"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
schematic_lemma tatsuta: "?p : (((P7 --> P1) --> P10) --> P4 --> P5)
--> (((P8 --> P2) --> P9) --> P3 --> P10)
--> (P1 --> P8) --> P6 --> P7
--> (((P3 --> P2) --> P9) --> P4)
--> (P1 --> P3) --> (((P6 --> P1) --> P2) --> P9) --> P5"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
schematic_lemma tatsuta1: "?p : (((P8 --> P2) --> P9) --> P3 --> P10)
--> (((P3 --> P2) --> P9) --> P4)
--> (((P6 --> P1) --> P2) --> P9)
--> (((P7 --> P1) --> P10) --> P4 --> P5)
--> (P1 --> P3) --> (P1 --> P8) --> P6 --> P7 --> P5"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
end
--- a/src/FOLP/ex/Quantifiers_Cla.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/FOLP/ex/Quantifiers_Cla.thy Thu Jul 23 14:25:05 2015 +0200
@@ -11,91 +11,91 @@
begin
schematic_lemma "?p : (ALL x y. P(x,y)) --> (ALL y x. P(x,y))"
- by (tactic {* Cla.fast_tac @{context} FOLP_cs 1 *})
+ by (tactic \<open>Cla.fast_tac @{context} FOLP_cs 1\<close>)
schematic_lemma "?p : (EX x y. P(x,y)) --> (EX y x. P(x,y))"
- by (tactic {* Cla.fast_tac @{context} FOLP_cs 1 *})
+ by (tactic \<open>Cla.fast_tac @{context} FOLP_cs 1\<close>)
(*Converse is false*)
schematic_lemma "?p : (ALL x. P(x)) | (ALL x. Q(x)) --> (ALL x. P(x) | Q(x))"
- by (tactic {* Cla.fast_tac @{context} FOLP_cs 1 *})
+ by (tactic \<open>Cla.fast_tac @{context} FOLP_cs 1\<close>)
schematic_lemma "?p : (ALL x. P-->Q(x)) <-> (P--> (ALL x. Q(x)))"
- by (tactic {* Cla.fast_tac @{context} FOLP_cs 1 *})
+ by (tactic \<open>Cla.fast_tac @{context} FOLP_cs 1\<close>)
schematic_lemma "?p : (ALL x. P(x)-->Q) <-> ((EX x. P(x)) --> Q)"
- by (tactic {* Cla.fast_tac @{context} FOLP_cs 1 *})
+ by (tactic \<open>Cla.fast_tac @{context} FOLP_cs 1\<close>)
text "Some harder ones"
schematic_lemma "?p : (EX x. P(x) | Q(x)) <-> (EX x. P(x)) | (EX x. Q(x))"
- by (tactic {* Cla.fast_tac @{context} FOLP_cs 1 *})
+ by (tactic \<open>Cla.fast_tac @{context} FOLP_cs 1\<close>)
(*Converse is false*)
schematic_lemma "?p : (EX x. P(x)&Q(x)) --> (EX x. P(x)) & (EX x. Q(x))"
- by (tactic {* Cla.fast_tac @{context} FOLP_cs 1 *})
+ by (tactic \<open>Cla.fast_tac @{context} FOLP_cs 1\<close>)
text "Basic test of quantifier reasoning"
(*TRUE*)
schematic_lemma "?p : (EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))"
- by (tactic {* Cla.fast_tac @{context} FOLP_cs 1 *})
+ by (tactic \<open>Cla.fast_tac @{context} FOLP_cs 1\<close>)
schematic_lemma "?p : (ALL x. Q(x)) --> (EX x. Q(x))"
- by (tactic {* Cla.fast_tac @{context} FOLP_cs 1 *})
+ by (tactic \<open>Cla.fast_tac @{context} FOLP_cs 1\<close>)
text "The following should fail, as they are false!"
schematic_lemma "?p : (ALL x. EX y. Q(x,y)) --> (EX y. ALL x. Q(x,y))"
- apply (tactic {* Cla.fast_tac @{context} FOLP_cs 1 *})?
+ apply (tactic \<open>Cla.fast_tac @{context} FOLP_cs 1\<close>)?
oops
schematic_lemma "?p : (EX x. Q(x)) --> (ALL x. Q(x))"
- apply (tactic {* Cla.fast_tac @{context} FOLP_cs 1 *})?
+ apply (tactic \<open>Cla.fast_tac @{context} FOLP_cs 1\<close>)?
oops
schematic_lemma "?p : P(?a) --> (ALL x. P(x))"
- apply (tactic {* Cla.fast_tac @{context} FOLP_cs 1 *})?
+ apply (tactic \<open>Cla.fast_tac @{context} FOLP_cs 1\<close>)?
oops
schematic_lemma "?p : (P(?a) --> (ALL x. Q(x))) --> (ALL x. P(x) --> Q(x))"
- apply (tactic {* Cla.fast_tac @{context} FOLP_cs 1 *})?
+ apply (tactic \<open>Cla.fast_tac @{context} FOLP_cs 1\<close>)?
oops
text "Back to things that are provable..."
schematic_lemma "?p : (ALL x. P(x)-->Q(x)) & (EX x. P(x)) --> (EX x. Q(x))"
- by (tactic {* Cla.fast_tac @{context} FOLP_cs 1 *})
+ by (tactic \<open>Cla.fast_tac @{context} FOLP_cs 1\<close>)
(*An example of why exI should be delayed as long as possible*)
schematic_lemma "?p : (P --> (EX x. Q(x))) & P --> (EX x. Q(x))"
- by (tactic {* Cla.fast_tac @{context} FOLP_cs 1 *})
+ by (tactic \<open>Cla.fast_tac @{context} FOLP_cs 1\<close>)
schematic_lemma "?p : (ALL x. P(x)-->Q(f(x))) & (ALL x. Q(x)-->R(g(x))) & P(d) --> R(?a)"
- by (tactic {* Cla.fast_tac @{context} FOLP_cs 1 *})
+ by (tactic \<open>Cla.fast_tac @{context} FOLP_cs 1\<close>)
schematic_lemma "?p : (ALL x. Q(x)) --> (EX x. Q(x))"
- by (tactic {* Cla.fast_tac @{context} FOLP_cs 1 *})
+ by (tactic \<open>Cla.fast_tac @{context} FOLP_cs 1\<close>)
text "Some slow ones"
(*Principia Mathematica *11.53 *)
schematic_lemma "?p : (ALL x y. P(x) --> Q(y)) <-> ((EX x. P(x)) --> (ALL y. Q(y)))"
- by (tactic {* Cla.fast_tac @{context} FOLP_cs 1 *})
+ by (tactic \<open>Cla.fast_tac @{context} FOLP_cs 1\<close>)
(*Principia Mathematica *11.55 *)
schematic_lemma "?p : (EX x y. P(x) & Q(x,y)) <-> (EX x. P(x) & (EX y. Q(x,y)))"
- by (tactic {* Cla.fast_tac @{context} FOLP_cs 1 *})
+ by (tactic \<open>Cla.fast_tac @{context} FOLP_cs 1\<close>)
(*Principia Mathematica *11.61 *)
schematic_lemma "?p : (EX y. ALL x. P(x) --> Q(x,y)) --> (ALL x. P(x) --> (EX y. Q(x,y)))"
- by (tactic {* Cla.fast_tac @{context} FOLP_cs 1 *})
+ by (tactic \<open>Cla.fast_tac @{context} FOLP_cs 1\<close>)
end
--- a/src/FOLP/ex/Quantifiers_Int.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/FOLP/ex/Quantifiers_Int.thy Thu Jul 23 14:25:05 2015 +0200
@@ -11,91 +11,91 @@
begin
schematic_lemma "?p : (ALL x y. P(x,y)) --> (ALL y x. P(x,y))"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
schematic_lemma "?p : (EX x y. P(x,y)) --> (EX y x. P(x,y))"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
(*Converse is false*)
schematic_lemma "?p : (ALL x. P(x)) | (ALL x. Q(x)) --> (ALL x. P(x) | Q(x))"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
schematic_lemma "?p : (ALL x. P-->Q(x)) <-> (P--> (ALL x. Q(x)))"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
schematic_lemma "?p : (ALL x. P(x)-->Q) <-> ((EX x. P(x)) --> Q)"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text "Some harder ones"
schematic_lemma "?p : (EX x. P(x) | Q(x)) <-> (EX x. P(x)) | (EX x. Q(x))"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
(*Converse is false*)
schematic_lemma "?p : (EX x. P(x)&Q(x)) --> (EX x. P(x)) & (EX x. Q(x))"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text "Basic test of quantifier reasoning"
(*TRUE*)
schematic_lemma "?p : (EX y. ALL x. Q(x,y)) --> (ALL x. EX y. Q(x,y))"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
schematic_lemma "?p : (ALL x. Q(x)) --> (EX x. Q(x))"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text "The following should fail, as they are false!"
schematic_lemma "?p : (ALL x. EX y. Q(x,y)) --> (EX y. ALL x. Q(x,y))"
- apply (tactic {* IntPr.fast_tac @{context} 1 *})?
+ apply (tactic \<open>IntPr.fast_tac @{context} 1\<close>)?
oops
schematic_lemma "?p : (EX x. Q(x)) --> (ALL x. Q(x))"
- apply (tactic {* IntPr.fast_tac @{context} 1 *})?
+ apply (tactic \<open>IntPr.fast_tac @{context} 1\<close>)?
oops
schematic_lemma "?p : P(?a) --> (ALL x. P(x))"
- apply (tactic {* IntPr.fast_tac @{context} 1 *})?
+ apply (tactic \<open>IntPr.fast_tac @{context} 1\<close>)?
oops
schematic_lemma "?p : (P(?a) --> (ALL x. Q(x))) --> (ALL x. P(x) --> Q(x))"
- apply (tactic {* IntPr.fast_tac @{context} 1 *})?
+ apply (tactic \<open>IntPr.fast_tac @{context} 1\<close>)?
oops
text "Back to things that are provable..."
schematic_lemma "?p : (ALL x. P(x)-->Q(x)) & (EX x. P(x)) --> (EX x. Q(x))"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
(*An example of why exI should be delayed as long as possible*)
schematic_lemma "?p : (P --> (EX x. Q(x))) & P --> (EX x. Q(x))"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
schematic_lemma "?p : (ALL x. P(x)-->Q(f(x))) & (ALL x. Q(x)-->R(g(x))) & P(d) --> R(?a)"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
schematic_lemma "?p : (ALL x. Q(x)) --> (EX x. Q(x))"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
text "Some slow ones"
(*Principia Mathematica *11.53 *)
schematic_lemma "?p : (ALL x y. P(x) --> Q(y)) <-> ((EX x. P(x)) --> (ALL y. Q(y)))"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
(*Principia Mathematica *11.55 *)
schematic_lemma "?p : (EX x y. P(x) & Q(x,y)) <-> (EX x. P(x) & (EX y. Q(x,y)))"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
(*Principia Mathematica *11.61 *)
schematic_lemma "?p : (EX y. ALL x. P(x) --> Q(x,y)) --> (ALL x. P(x) --> (EX y. Q(x,y)))"
- by (tactic {* IntPr.fast_tac @{context} 1 *})
+ by (tactic \<open>IntPr.fast_tac @{context} 1\<close>)
end
--- a/src/LCF/LCF.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/LCF/LCF.thy Thu Jul 23 14:25:05 2015 +0200
@@ -3,15 +3,15 @@
Copyright 1992 University of Cambridge
*)
-section {* LCF on top of First-Order Logic *}
+section \<open>LCF on top of First-Order Logic\<close>
theory LCF
imports "~~/src/FOL/FOL"
begin
-text {* This theory is based on Lawrence Paulson's book Logic and Computation. *}
+text \<open>This theory is based on Lawrence Paulson's book Logic and Computation.\<close>
-subsection {* Natural Deduction Rules for LCF *}
+subsection \<open>Natural Deduction Rules for LCF\<close>
class cpo = "term"
default_sort cpo
@@ -245,7 +245,7 @@
surj_pairing FST SND
-subsection {* Ordered pairs and products *}
+subsection \<open>Ordered pairs and products\<close>
lemma expand_all_PROD: "(\<forall>p. P(p)) \<longleftrightarrow> (\<forall>x y. P(<x,y>))"
apply (rule iffI)
@@ -293,7 +293,7 @@
done
-subsection {* Fixedpoint theory *}
+subsection \<open>Fixedpoint theory\<close>
lemma adm_eq: "adm(\<lambda>x. t(x)=(u(x)::'a::cpo))"
apply (unfold eq_def)
@@ -318,12 +318,12 @@
adm_not_free adm_eq adm_less adm_not_less
adm_not_eq_tr adm_conj adm_disj adm_imp adm_all
-method_setup induct = {*
+method_setup induct = \<open>
Scan.lift Args.name_inner_syntax >> (fn v => fn ctxt =>
SIMPLE_METHOD' (fn i =>
Rule_Insts.res_inst_tac ctxt [((("f", 0), Position.none), v)] [] @{thm induct} i THEN
REPEAT (resolve_tac ctxt @{thms adm_lemmas} i)))
-*}
+\<close>
lemma least_FIX: "f(p) = p \<Longrightarrow> FIX(f) << p"
apply (induct f)
@@ -378,11 +378,11 @@
apply (rule 3)
done
-ML {*
+ML \<open>
fun induct2_tac ctxt (f, g) i =
Rule_Insts.res_inst_tac ctxt
[((("f", 0), Position.none), f), ((("g", 0), Position.none), g)] [] @{thm induct2} i THEN
REPEAT(resolve_tac ctxt @{thms adm_lemmas} i)
-*}
+\<close>
end
--- a/src/LCF/ex/Ex1.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/LCF/ex/Ex1.thy Thu Jul 23 14:25:05 2015 +0200
@@ -1,4 +1,4 @@
-section {* Section 10.4 *}
+section \<open>Section 10.4\<close>
theory Ex1
imports "../LCF"
--- a/src/LCF/ex/Ex2.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/LCF/ex/Ex2.thy Thu Jul 23 14:25:05 2015 +0200
@@ -1,4 +1,4 @@
-section {* Example 3.8 *}
+section \<open>Example 3.8\<close>
theory Ex2
imports "../LCF"
--- a/src/LCF/ex/Ex3.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/LCF/ex/Ex3.thy Thu Jul 23 14:25:05 2015 +0200
@@ -1,4 +1,4 @@
-section {* Addition with fixpoint of successor *}
+section \<open>Addition with fixpoint of successor\<close>
theory Ex3
imports "../LCF"
--- a/src/LCF/ex/Ex4.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/LCF/ex/Ex4.thy Thu Jul 23 14:25:05 2015 +0200
@@ -1,5 +1,5 @@
-section {* Prefixpoints *}
+section \<open>Prefixpoints\<close>
theory Ex4
imports "../LCF"
--- a/src/Sequents/ILL.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/Sequents/ILL.thy Thu Jul 23 14:25:05 2015 +0200
@@ -35,17 +35,17 @@
"_Context" :: "two_seqe" ("((_)/ :=: (_))" [6,6] 5)
"_PromAux" :: "three_seqe" ("promaux {_||_||_}")
-parse_translation {*
+parse_translation \<open>
[(@{syntax_const "_Trueprop"}, K (single_tr @{const_syntax Trueprop})),
(@{syntax_const "_Context"}, K (two_seq_tr @{const_syntax Context})),
(@{syntax_const "_PromAux"}, K (three_seq_tr @{const_syntax PromAux}))]
-*}
+\<close>
-print_translation {*
+print_translation \<open>
[(@{const_syntax Trueprop}, K (single_tr' @{syntax_const "_Trueprop"})),
(@{const_syntax Context}, K (two_seq_tr' @{syntax_const "_Context"})),
(@{const_syntax PromAux}, K (three_seq_tr' @{syntax_const "_PromAux"}))]
-*}
+\<close>
defs
@@ -271,7 +271,7 @@
apply best
done
-ML {*
+ML \<open>
val safe_pack =
@{context}
|> fold_rev Cla.add_safe @{thms conj_lemma ll_mp contrad1
@@ -283,13 +283,13 @@
Cla.put_pack safe_pack @{context}
|> Cla.add_unsafe @{thm impr_contr_der}
|> Cla.get_pack;
-*}
+\<close>
method_setup best_safe =
- {* Scan.succeed (fn ctxt => SIMPLE_METHOD' (Cla.best_tac (Cla.put_pack safe_pack ctxt))) *}
+ \<open>Scan.succeed (fn ctxt => SIMPLE_METHOD' (Cla.best_tac (Cla.put_pack safe_pack ctxt)))\<close>
method_setup best_power =
- {* Scan.succeed (fn ctxt => SIMPLE_METHOD' (Cla.best_tac (Cla.put_pack power_pack ctxt))) *}
+ \<open>Scan.succeed (fn ctxt => SIMPLE_METHOD' (Cla.best_tac (Cla.put_pack power_pack ctxt)))\<close>
(* Some examples from Troelstra and van Dalen *)
--- a/src/Sequents/LK.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/Sequents/LK.thy Thu Jul 23 14:25:05 2015 +0200
@@ -23,7 +23,7 @@
==> (P, $H |- $F) == (P', $H' |- $F')"
-subsection {* Rewrite rules *}
+subsection \<open>Rewrite rules\<close>
lemma conj_simps:
"|- P & True <-> P"
@@ -79,15 +79,15 @@
by (fast add!: subst)+
-subsection {* Miniscoping: pushing quantifiers in *}
+subsection \<open>Miniscoping: pushing quantifiers in\<close>
-text {*
+text \<open>
We do NOT distribute of ALL over &, or dually that of EX over |
Baaz and Leitsch, On Skolemization and Proof Complexity (1994)
show that this step can increase proof length!
-*}
+\<close>
-text {*existential miniscoping*}
+text \<open>existential miniscoping\<close>
lemma ex_simps:
"!!P Q. |- (EX x. P(x) & Q) <-> (EX x. P(x)) & Q"
"!!P Q. |- (EX x. P & Q(x)) <-> P & (EX x. Q(x))"
@@ -97,7 +97,7 @@
"!!P Q. |- (EX x. P --> Q(x)) <-> P --> (EX x. Q(x))"
by (fast add!: subst)+
-text {*universal miniscoping*}
+text \<open>universal miniscoping\<close>
lemma all_simps:
"!!P Q. |- (ALL x. P(x) & Q) <-> (ALL x. P(x)) & Q"
"!!P Q. |- (ALL x. P & Q(x)) <-> P & (ALL x. Q(x))"
@@ -107,7 +107,7 @@
"!!P Q. |- (ALL x. P | Q(x)) <-> P | (ALL x. Q(x))"
by (fast add!: subst)+
-text {*These are NOT supplied by default!*}
+text \<open>These are NOT supplied by default!\<close>
lemma distrib_simps:
"|- P & (Q | R) <-> P&Q | P&R"
"|- (Q | R) & P <-> Q&P | R&P"
@@ -138,7 +138,7 @@
by (fast add!: subst)+
-subsection {* Named rewrite rules *}
+subsection \<open>Named rewrite rules\<close>
lemma conj_commute: "|- P&Q <-> Q&P"
and conj_left_commute: "|- P&(Q&R) <-> Q&(P&R)"
@@ -177,11 +177,11 @@
shows "|- (P-->Q) <-> (P'-->Q')"
apply (lem p1)
apply safe
- apply (tactic {*
+ apply (tactic \<open>
REPEAT (resolve_tac @{context} @{thms cut} 1 THEN
DEPTH_SOLVE_1
(resolve_tac @{context} [@{thm thinL}, @{thm thinR}, @{thm p2} COMP @{thm monotonic}] 1) THEN
- Cla.safe_tac @{context} 1) *})
+ Cla.safe_tac @{context} 1)\<close>)
done
lemma conj_cong:
@@ -190,22 +190,22 @@
shows "|- (P&Q) <-> (P'&Q')"
apply (lem p1)
apply safe
- apply (tactic {*
+ apply (tactic \<open>
REPEAT (resolve_tac @{context} @{thms cut} 1 THEN
DEPTH_SOLVE_1
(resolve_tac @{context} [@{thm thinL}, @{thm thinR}, @{thm p2} COMP @{thm monotonic}] 1) THEN
- Cla.safe_tac @{context} 1) *})
+ Cla.safe_tac @{context} 1)\<close>)
done
lemma eq_sym_conv: "|- (x=y) <-> (y=x)"
by (fast add!: subst)
ML_file "simpdata.ML"
-setup {* map_theory_simpset (put_simpset LK_ss) *}
-setup {* Simplifier.method_setup [] *}
+setup \<open>map_theory_simpset (put_simpset LK_ss)\<close>
+setup \<open>Simplifier.method_setup []\<close>
-text {* To create substition rules *}
+text \<open>To create substition rules\<close>
lemma eq_imp_subst: "|- a=b ==> $H, A(a), $G |- $E, A(b), $F"
by simp
--- a/src/Sequents/LK/Nat.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/Sequents/LK/Nat.thy Thu Jul 23 14:25:05 2015 +0200
@@ -3,7 +3,7 @@
Copyright 1999 University of Cambridge
*)
-section {* Theory of the natural numbers: Peano's axioms, primitive recursion *}
+section \<open>Theory of the natural numbers: Peano's axioms, primitive recursion\<close>
theory Nat
imports "../LK"
--- a/src/Sequents/LK/Propositional.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/Sequents/LK/Propositional.thy Thu Jul 23 14:25:05 2015 +0200
@@ -3,7 +3,7 @@
Copyright 1992 University of Cambridge
*)
-section {* Classical sequent calculus: examples with propositional connectives *}
+section \<open>Classical sequent calculus: examples with propositional connectives\<close>
theory Propositional
imports "../LK"
--- a/src/Sequents/LK0.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/Sequents/LK0.thy Thu Jul 23 14:25:05 2015 +0200
@@ -6,7 +6,7 @@
(eta-expanded, beta-contracted).
*)
-section {* Classical First-Order Sequent Calculus *}
+section \<open>Classical First-Order Sequent Calculus\<close>
theory LK0
imports Sequents
@@ -34,8 +34,8 @@
syntax
"_Trueprop" :: "two_seqe" ("((_)/ |- (_))" [6,6] 5)
-parse_translation {* [(@{syntax_const "_Trueprop"}, K (two_seq_tr @{const_syntax Trueprop}))] *}
-print_translation {* [(@{const_syntax Trueprop}, K (two_seq_tr' @{syntax_const "_Trueprop"}))] *}
+parse_translation \<open>[(@{syntax_const "_Trueprop"}, K (two_seq_tr @{const_syntax Trueprop}))]\<close>
+print_translation \<open>[(@{const_syntax Trueprop}, K (two_seq_tr' @{syntax_const "_Trueprop"}))]\<close>
abbreviation
not_equal (infixl "~=" 50) where
@@ -150,7 +150,7 @@
lemma exchL: "$H, Q, P, $G |- $E ==> $H, P, Q, $G |- $E"
by (rule exchLS)
-ML {*
+ML \<open>
(*Cut and thin, replacing the right-side formula*)
fun cutR_tac ctxt s i =
Rule_Insts.res_inst_tac ctxt [((("P", 0), Position.none), s)] [] @{thm cut} i THEN
@@ -160,7 +160,7 @@
fun cutL_tac ctxt s i =
Rule_Insts.res_inst_tac ctxt [((("P", 0), Position.none), s)] [] @{thm cut} i THEN
resolve_tac ctxt @{thms thinL} (i + 1)
-*}
+\<close>
(** If-and-only-if rules **)
@@ -220,36 +220,36 @@
conjR conjL
FalseL TrueR
refl basic
-ML {* val prop_pack = Cla.get_pack @{context} *}
+ML \<open>val prop_pack = Cla.get_pack @{context}\<close>
lemmas [safe] = exL allR
lemmas [unsafe] = the_equality exR_thin allL_thin
-ML {* val LK_pack = Cla.get_pack @{context} *}
+ML \<open>val LK_pack = Cla.get_pack @{context}\<close>
-ML {*
+ML \<open>
val LK_dup_pack =
Cla.put_pack prop_pack @{context}
|> fold_rev Cla.add_safe @{thms allR exL}
|> fold_rev Cla.add_unsafe @{thms allL exR the_equality}
|> Cla.get_pack;
-*}
+\<close>
method_setup fast_prop =
- {* Scan.succeed (fn ctxt => SIMPLE_METHOD' (Cla.fast_tac (Cla.put_pack prop_pack ctxt))) *}
+ \<open>Scan.succeed (fn ctxt => SIMPLE_METHOD' (Cla.fast_tac (Cla.put_pack prop_pack ctxt)))\<close>
method_setup fast_dup =
- {* Scan.succeed (fn ctxt => SIMPLE_METHOD' (Cla.fast_tac (Cla.put_pack LK_dup_pack ctxt))) *}
+ \<open>Scan.succeed (fn ctxt => SIMPLE_METHOD' (Cla.fast_tac (Cla.put_pack LK_dup_pack ctxt)))\<close>
method_setup best_dup =
- {* Scan.succeed (fn ctxt => SIMPLE_METHOD' (Cla.best_tac (Cla.put_pack LK_dup_pack ctxt))) *}
+ \<open>Scan.succeed (fn ctxt => SIMPLE_METHOD' (Cla.best_tac (Cla.put_pack LK_dup_pack ctxt)))\<close>
-method_setup lem = {*
+method_setup lem = \<open>
Attrib.thm >> (fn th => fn ctxt =>
SIMPLE_METHOD' (fn i =>
resolve_tac ctxt [@{thm thinR} RS @{thm cut}] i THEN
REPEAT (resolve_tac ctxt @{thms thinL} i) THEN
resolve_tac ctxt [th] i))
-*}
+\<close>
lemma mp_R:
--- a/src/Sequents/Modal0.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/Sequents/Modal0.thy Thu Jul 23 14:25:05 2015 +0200
@@ -21,20 +21,20 @@
"_Lstar" :: "two_seqe" ("(_)|L>(_)" [6,6] 5)
"_Rstar" :: "two_seqe" ("(_)|R>(_)" [6,6] 5)
-ML {*
+ML \<open>
fun star_tr c [s1, s2] = Const(c, dummyT) $ seq_tr s1 $ seq_tr s2;
fun star_tr' c [s1, s2] = Const(c, dummyT) $ seq_tr' s1 $ seq_tr' s2;
-*}
+\<close>
-parse_translation {*
+parse_translation \<open>
[(@{syntax_const "_Lstar"}, K (star_tr @{const_syntax Lstar})),
(@{syntax_const "_Rstar"}, K (star_tr @{const_syntax Rstar}))]
-*}
+\<close>
-print_translation {*
+print_translation \<open>
[(@{const_syntax Lstar}, K (star_tr' @{syntax_const "_Lstar"})),
(@{const_syntax Rstar}, K (star_tr' @{syntax_const "_Rstar"}))]
-*}
+\<close>
defs
strimp_def: "P --< Q == [](P --> Q)"
--- a/src/Sequents/S4.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/Sequents/S4.thy Thu Jul 23 14:25:05 2015 +0200
@@ -31,7 +31,7 @@
"[| $E |L> $E'; $F |L> $F'; $G |R> $G';
$E', P, $F' |- $G'|] ==> $E, <>P, $F |- $G"
-ML {*
+ML \<open>
structure S4_Prover = Modal_ProverFun
(
val rewrite_rls = @{thms rewrite_rls}
@@ -41,10 +41,10 @@
val aside_rls = [@{thm lstar0}, @{thm lstar1}, @{thm lstar2}, @{thm rstar0},
@{thm rstar1}, @{thm rstar2}]
)
-*}
+\<close>
method_setup S4_solve =
- {* Scan.succeed (fn ctxt => SIMPLE_METHOD (S4_Prover.solve_tac ctxt 2)) *}
+ \<open>Scan.succeed (fn ctxt => SIMPLE_METHOD (S4_Prover.solve_tac ctxt 2))\<close>
(* Theorems of system T from Hughes and Cresswell and Hailpern, LNCS 129 *)
--- a/src/Sequents/S43.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/Sequents/S43.thy Thu Jul 23 14:25:05 2015 +0200
@@ -16,21 +16,21 @@
"_S43pi" :: "[seq, seq, seq, seq, seq, seq] => prop"
("S43pi((_);(_);(_);(_);(_);(_))" [] 5)
-parse_translation {*
+parse_translation \<open>
let
val tr = seq_tr;
fun s43pi_tr [s1, s2, s3, s4, s5, s6] =
Const (@{const_syntax S43pi}, dummyT) $ tr s1 $ tr s2 $ tr s3 $ tr s4 $ tr s5 $ tr s6;
in [(@{syntax_const "_S43pi"}, K s43pi_tr)] end
-*}
+\<close>
-print_translation {*
+print_translation \<open>
let
val tr' = seq_tr';
fun s43pi_tr' [s1, s2, s3, s4, s5, s6] =
Const(@{syntax_const "_S43pi"}, dummyT) $ tr' s1 $ tr' s2 $ tr' s3 $ tr' s4 $ tr' s5 $ tr' s6;
in [(@{const_syntax S43pi}, K s43pi_tr')] end
-*}
+\<close>
axiomatization where
(* Definition of the star operation using a set of Horn clauses *)
@@ -76,7 +76,7 @@
$L |- $R1, []P, $R2"
-ML {*
+ML \<open>
structure S43_Prover = Modal_ProverFun
(
val rewrite_rls = @{thms rewrite_rls}
@@ -86,13 +86,13 @@
val aside_rls = [@{thm lstar0}, @{thm lstar1}, @{thm lstar2}, @{thm rstar0},
@{thm rstar1}, @{thm rstar2}, @{thm S43pi0}, @{thm S43pi1}, @{thm S43pi2}]
)
-*}
+\<close>
-method_setup S43_solve = {*
+method_setup S43_solve = \<open>
Scan.succeed (fn ctxt => SIMPLE_METHOD
(S43_Prover.solve_tac ctxt 2 ORELSE S43_Prover.solve_tac ctxt 3))
-*}
+\<close>
(* Theorems of system T from Hughes and Cresswell and Hailpern, LNCS 129 *)
--- a/src/Sequents/Sequents.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/Sequents/Sequents.thy Thu Jul 23 14:25:05 2015 +0200
@@ -3,7 +3,7 @@
Copyright 1993 University of Cambridge
*)
-section {* Parsing and pretty-printing of sequences *}
+section \<open>Parsing and pretty-printing of sequences\<close>
theory Sequents
imports Pure
@@ -17,7 +17,7 @@
typedecl o
-subsection {* Sequences *}
+subsection \<open>Sequences\<close>
typedecl
seq'
@@ -27,7 +27,7 @@
Seq1' :: "o=>seq'"
-subsection {* Concrete syntax *}
+subsection \<open>Concrete syntax\<close>
nonterminal seq and seqobj and seqcont
@@ -56,7 +56,7 @@
(*Constant to allow definitions of SEQUENCES of formulas*)
"_Side" :: "seq=>(seq'=>seq')" ("<<(_)>>")
-ML {*
+ML \<open>
(* parse translation for sequences *)
@@ -139,12 +139,12 @@
(** for the <<...>> notation **)
fun side_tr [s1] = seq_tr s1;
-*}
+\<close>
-parse_translation {* [(@{syntax_const "_Side"}, K side_tr)] *}
+parse_translation \<open>[(@{syntax_const "_Side"}, K side_tr)]\<close>
-subsection {* Proof tools *}
+subsection \<open>Proof tools\<close>
ML_file "prover.ML"
--- a/src/Sequents/T.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/Sequents/T.thy Thu Jul 23 14:25:05 2015 +0200
@@ -30,7 +30,7 @@
"[| $E |L> $E'; $F |L> $F'; $G |R> $G';
$E', P, $F'|- $G'|] ==> $E, <>P, $F |- $G"
-ML {*
+ML \<open>
structure T_Prover = Modal_ProverFun
(
val rewrite_rls = @{thms rewrite_rls}
@@ -40,9 +40,9 @@
val aside_rls = [@{thm lstar0}, @{thm lstar1}, @{thm lstar2}, @{thm rstar0},
@{thm rstar1}, @{thm rstar2}]
)
-*}
+\<close>
-method_setup T_solve = {* Scan.succeed (fn ctxt => SIMPLE_METHOD (T_Prover.solve_tac ctxt 2)) *}
+method_setup T_solve = \<open>Scan.succeed (fn ctxt => SIMPLE_METHOD (T_Prover.solve_tac ctxt 2))\<close>
(* Theorems of system T from Hughes and Cresswell and Hailpern, LNCS 129 *)
--- a/src/Sequents/Washing.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/Sequents/Washing.thy Thu Jul 23 14:25:05 2015 +0200
@@ -32,10 +32,10 @@
(* "activate" definitions for use in proof *)
-ML {* ML_Thms.bind_thms ("changeI", [@{thm context1}] RL ([@{thm change}] RLN (2,[@{thm cut}]))) *}
-ML {* ML_Thms.bind_thms ("load1I", [@{thm context1}] RL ([@{thm load1}] RLN (2,[@{thm cut}]))) *}
-ML {* ML_Thms.bind_thms ("washI", [@{thm context1}] RL ([@{thm wash}] RLN (2,[@{thm cut}]))) *}
-ML {* ML_Thms.bind_thms ("dryI", [@{thm context1}] RL ([@{thm dry}] RLN (2,[@{thm cut}]))) *}
+ML \<open>ML_Thms.bind_thms ("changeI", [@{thm context1}] RL ([@{thm change}] RLN (2,[@{thm cut}])))\<close>
+ML \<open>ML_Thms.bind_thms ("load1I", [@{thm context1}] RL ([@{thm load1}] RLN (2,[@{thm cut}])))\<close>
+ML \<open>ML_Thms.bind_thms ("washI", [@{thm context1}] RL ([@{thm wash}] RLN (2,[@{thm cut}])))\<close>
+ML \<open>ML_Thms.bind_thms ("dryI", [@{thm context1}] RL ([@{thm dry}] RLN (2,[@{thm cut}])))\<close>
(* a load of dirty clothes and two dollars gives you clean clothes *)
--- a/src/ZF/AC.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/ZF/AC.thy Thu Jul 23 14:25:05 2015 +0200
@@ -3,11 +3,11 @@
Copyright 1994 University of Cambridge
*)
-section{*The Axiom of Choice*}
+section\<open>The Axiom of Choice\<close>
theory AC imports Main_ZF begin
-text{*This definition comes from Halmos (1960), page 59.*}
+text\<open>This definition comes from Halmos (1960), page 59.\<close>
axiomatization where
AC: "[| a \<in> A; !!x. x \<in> A ==> (\<exists>y. y \<in> B(x)) |] ==> \<exists>z. z \<in> Pi(A,B)"
--- a/src/ZF/AC/Cardinal_aux.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/ZF/AC/Cardinal_aux.thy Thu Jul 23 14:25:05 2015 +0200
@@ -45,7 +45,7 @@
by (simp add: Inf_Card_is_InfCard Finite_cardinal_iff NFI i)
have "A \<union> B \<lesssim> A + B" by (rule Un_lepoll_sum)
also have "... \<lesssim> A \<times> B"
- by (rule lepoll_imp_sum_lepoll_prod [OF AB [THEN eqpoll_imp_lepoll] `2 \<lesssim> A`])
+ by (rule lepoll_imp_sum_lepoll_prod [OF AB [THEN eqpoll_imp_lepoll] \<open>2 \<lesssim> A\<close>])
also have "... \<approx> i \<times> i"
by (blast intro: prod_eqpoll_cong eqpoll_imp_lepoll A B)
also have "... \<approx> i"
@@ -172,11 +172,11 @@
apply (drule eqpoll_imp_lepoll [THEN lepoll_trans],
rule le_imp_lepoll, assumption)+
apply (case_tac "Finite(x \<union> xa)")
-txt{*finite case*}
+txt\<open>finite case\<close>
apply (drule Finite_Un [OF lepoll_Finite lepoll_Finite], assumption+)
apply (drule subset_Un_Diff [THEN subset_imp_lepoll, THEN lepoll_Finite])
apply (fast dest: eqpoll_sym [THEN eqpoll_imp_lepoll, THEN lepoll_Finite])
-txt{*infinite case*}
+txt\<open>infinite case\<close>
apply (drule Un_lepoll_Inf_Ord, (assumption+))
apply (blast intro: le_Ord2)
apply (drule lesspoll_trans1
--- a/src/ZF/AC/DC.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/ZF/AC/DC.thy Thu Jul 23 14:25:05 2015 +0200
@@ -17,7 +17,7 @@
apply (fast intro: LeastI Ord_in_Ord)
done
-text{*Trivial in the presence of AC, but here we need a wellordering of X*}
+text\<open>Trivial in the presence of AC, but here we need a wellordering of X\<close>
lemma image_Ord_lepoll: "[| f \<in> X->Y; Ord(X) |] ==> f``X \<lesssim> X"
apply (unfold lepoll_def)
apply (rule_tac x = "\<lambda>x \<in> f``X. LEAST y. f`y = x" in exI)
--- a/src/ZF/AC/HH.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/ZF/AC/HH.thy Thu Jul 23 14:25:05 2015 +0200
@@ -16,7 +16,7 @@
"HH(f,x,a) == transrec(a, %b r. let z = x - (\<Union>c \<in> b. r`c)
in if f`z \<in> Pow(z)-{0} then f`z else {x})"
-subsection{*Lemmas useful in each of the three proofs*}
+subsection\<open>Lemmas useful in each of the three proofs\<close>
lemma HH_def_satisfies_eq:
"HH(f,x,a) = (let z = x - (\<Union>b \<in> a. HH(f,x,b))
@@ -126,7 +126,7 @@
apply (erule_tac [2] ltI [OF _ Ord_Least], assumption)
done
-subsection{*Lemmas used in the proofs of AC1 ==> WO2 and AC17 ==> AC1*}
+subsection\<open>Lemmas used in the proofs of AC1 ==> WO2 and AC17 ==> AC1\<close>
lemma lam_Least_HH_inj_Pow:
"(\<lambda>a \<in> (LEAST i. HH(f,x,i)={x}). HH(f,x,a))
@@ -213,7 +213,7 @@
lam_sing_bij [THEN bij_converse_bij]]
-subsection{*The proof of AC1 ==> WO2*}
+subsection\<open>The proof of AC1 ==> WO2\<close>
(*Establishing the existence of a bijection, namely
converse
--- a/src/ZF/Arith.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/ZF/Arith.thy Thu Jul 23 14:25:05 2015 +0200
@@ -9,11 +9,11 @@
Also, rec(m, 0, %z w.z) is pred(m).
*)
-section{*Arithmetic Operators and Their Definitions*}
+section\<open>Arithmetic Operators and Their Definitions\<close>
theory Arith imports Univ begin
-text{*Proofs about elementary arithmetic: addition, multiplication, etc.*}
+text\<open>Proofs about elementary arithmetic: addition, multiplication, etc.\<close>
definition
pred :: "i=>i" (*inverse of succ*) where
@@ -91,7 +91,7 @@
lemmas zero_lt_natE = zero_lt_lemma [THEN bexE]
-subsection{*@{text natify}, the Coercion to @{term nat}*}
+subsection\<open>@{text natify}, the Coercion to @{term nat}\<close>
lemma pred_succ_eq [simp]: "pred(succ(y)) = y"
by (unfold pred_def, auto)
@@ -167,7 +167,7 @@
by (simp add: div_def)
-subsection{*Typing rules*}
+subsection\<open>Typing rules\<close>
(** Addition **)
@@ -222,7 +222,7 @@
done
-subsection{*Addition*}
+subsection\<open>Addition\<close>
(*Natify has weakened this law, compared with the older approach*)
lemma add_0_natify [simp]: "0 #+ m = natify(m)"
@@ -315,7 +315,7 @@
by (induct_tac "n", auto)
-subsection{*Monotonicity of Addition*}
+subsection\<open>Monotonicity of Addition\<close>
(*strict, in 1st argument; proof is by rule induction on 'less than'.
Still need j\<in>nat, for consider j = omega. Then we can have i<omega,
@@ -326,11 +326,11 @@
apply (simp_all add: leI)
done
-text{*strict, in second argument*}
+text\<open>strict, in second argument\<close>
lemma add_lt_mono2: "[| i<j; j\<in>nat |] ==> k#+i < k#+j"
by (simp add: add_commute [of k] add_lt_mono1)
-text{*A [clumsy] way of lifting < monotonicity to @{text "\<le>"} monotonicity*}
+text\<open>A [clumsy] way of lifting < monotonicity to @{text "\<le>"} monotonicity\<close>
lemma Ord_lt_mono_imp_le_mono:
assumes lt_mono: "!!i j. [| i<j; j:k |] ==> f(i) < f(j)"
and ford: "!!i. i:k ==> Ord(f(i))"
@@ -341,19 +341,19 @@
apply (blast intro!: leCI lt_mono ford elim!: leE)
done
-text{*@{text "\<le>"} monotonicity, 1st argument*}
+text\<open>@{text "\<le>"} monotonicity, 1st argument\<close>
lemma add_le_mono1: "[| i \<le> j; j\<in>nat |] ==> i#+k \<le> j#+k"
apply (rule_tac f = "%j. j#+k" in Ord_lt_mono_imp_le_mono, typecheck)
apply (blast intro: add_lt_mono1 add_type [THEN nat_into_Ord])+
done
-text{*@{text "\<le>"} monotonicity, both arguments*}
+text\<open>@{text "\<le>"} monotonicity, both arguments\<close>
lemma add_le_mono: "[| i \<le> j; k \<le> l; j\<in>nat; l\<in>nat |] ==> i#+k \<le> j#+l"
apply (rule add_le_mono1 [THEN le_trans], assumption+)
apply (subst add_commute, subst add_commute, rule add_le_mono1, assumption+)
done
-text{*Combinations of less-than and less-than-or-equals*}
+text\<open>Combinations of less-than and less-than-or-equals\<close>
lemma add_lt_le_mono: "[| i<j; k\<le>l; j\<in>nat; l\<in>nat |] ==> i#+k < j#+l"
apply (rule add_lt_mono1 [THEN lt_trans2], assumption+)
@@ -363,7 +363,7 @@
lemma add_le_lt_mono: "[| i\<le>j; k<l; j\<in>nat; l\<in>nat |] ==> i#+k < j#+l"
by (subst add_commute, subst add_commute, erule add_lt_le_mono, assumption+)
-text{*Less-than: in other words, strict in both arguments*}
+text\<open>Less-than: in other words, strict in both arguments\<close>
lemma add_lt_mono: "[| i<j; k<l; j\<in>nat; l\<in>nat |] ==> i#+k < j#+l"
apply (rule add_lt_le_mono)
apply (auto intro: leI)
@@ -433,7 +433,7 @@
"[|i\<in>nat; j\<in>nat|] ==> (i \<union> j) #- k = (i#-k) \<union> (j#-k)"
by (insert nat_diff_Un_distrib [of i j "natify(k)"], simp)
-text{*We actually prove @{term "i #- j #- k = i #- (j #+ k)"}*}
+text\<open>We actually prove @{term "i #- j #- k = i #- (j #+ k)"}\<close>
lemma diff_diff_left [simplified]:
"natify(i)#-natify(j)#-k = natify(i) #- (natify(j)#+k)"
by (rule_tac m="natify(i)" and n="natify(j)" in diff_induct, auto)
@@ -464,7 +464,7 @@
by auto
-subsection{*Multiplication*}
+subsection\<open>Multiplication\<close>
lemma mult_0 [simp]: "0 #* m = 0"
by (simp add: mult_def)
--- a/src/ZF/ArithSimp.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/ZF/ArithSimp.thy Thu Jul 23 14:25:05 2015 +0200
@@ -3,7 +3,7 @@
Copyright 2000 University of Cambridge
*)
-section{*Arithmetic with simplification*}
+section\<open>Arithmetic with simplification\<close>
theory ArithSimp
imports Arith
@@ -14,7 +14,7 @@
ML_file "arith_data.ML"
-subsection{*Difference*}
+subsection\<open>Difference\<close>
lemma diff_self_eq_0 [simp]: "m #- m = 0"
apply (subgoal_tac "natify (m) #- natify (m) = 0")
@@ -64,7 +64,7 @@
done
-subsection{*Remainder*}
+subsection\<open>Remainder\<close>
(*We need m:nat even with natify*)
lemma div_termination: "[| 0<n; n \<le> m; m:nat |] ==> m #- n < m"
@@ -135,7 +135,7 @@
done
-subsection{*Division*}
+subsection\<open>Division\<close>
lemma raw_div_type: "[| m:nat; n:nat |] ==> raw_div (m, n) \<in> nat"
apply (unfold raw_div_def)
@@ -183,9 +183,9 @@
apply (simp add: nat_into_Ord [THEN Ord_0_lt_iff])
apply (erule complete_induct)
apply (case_tac "x<n")
-txt{*case x<n*}
+txt\<open>case x<n\<close>
apply (simp (no_asm_simp))
-txt{*case @{term"n \<le> x"}*}
+txt\<open>case @{term"n \<le> x"}\<close>
apply (simp add: not_lt_iff_le add_assoc mod_geq div_termination [THEN ltD] add_diff_inverse)
done
@@ -200,23 +200,23 @@
done
-subsection{*Further Facts about Remainder*}
+subsection\<open>Further Facts about Remainder\<close>
-text{*(mainly for mutilated chess board)*}
+text\<open>(mainly for mutilated chess board)\<close>
lemma mod_succ_lemma:
"[| 0<n; m:nat; n:nat |]
==> succ(m) mod n = (if succ(m mod n) = n then 0 else succ(m mod n))"
apply (erule complete_induct)
apply (case_tac "succ (x) <n")
-txt{* case succ(x) < n *}
+txt\<open>case succ(x) < n\<close>
apply (simp (no_asm_simp) add: nat_le_refl [THEN lt_trans] succ_neq_self)
apply (simp add: ltD [THEN mem_imp_not_eq])
-txt{* case @{term"n \<le> succ(x)"} *}
+txt\<open>case @{term"n \<le> succ(x)"}\<close>
apply (simp add: mod_geq not_lt_iff_le)
apply (erule leE)
apply (simp (no_asm_simp) add: mod_geq div_termination [THEN ltD] diff_succ)
-txt{*equality case*}
+txt\<open>equality case\<close>
apply (simp add: diff_self_eq_0)
done
@@ -235,7 +235,7 @@
apply (subgoal_tac "natify (m) mod n < n")
apply (rule_tac [2] i = "natify (m) " in complete_induct)
apply (case_tac [3] "x<n", auto)
-txt{* case @{term"n \<le> x"}*}
+txt\<open>case @{term"n \<le> x"}\<close>
apply (simp add: mod_geq not_lt_iff_le div_termination [THEN ltD])
done
@@ -264,7 +264,7 @@
by (cut_tac n = 0 in mod2_add_more, auto)
-subsection{*Additional theorems about @{text "\<le>"}*}
+subsection\<open>Additional theorems about @{text "\<le>"}\<close>
lemma add_le_self: "m:nat ==> m \<le> (m #+ n)"
apply (simp (no_asm_simp))
@@ -339,7 +339,7 @@
done
-subsection{*Cancellation Laws for Common Factors in Comparisons*}
+subsection\<open>Cancellation Laws for Common Factors in Comparisons\<close>
lemma mult_less_cancel_lemma:
"[| k: nat; m: nat; n: nat |] ==> (m#*k < n#*k) \<longleftrightarrow> (0<k & m<n)"
@@ -414,7 +414,7 @@
done
-subsection{*More Lemmas about Remainder*}
+subsection\<open>More Lemmas about Remainder\<close>
lemma mult_mod_distrib_raw:
"[| k:nat; m:nat; n:nat |] ==> (k#*m) mod (k#*n) = k #* (m mod n)"
@@ -507,7 +507,7 @@
by (drule less_imp_succ_add, auto)
-subsubsection{*More Lemmas About Difference*}
+subsubsection\<open>More Lemmas About Difference\<close>
lemma diff_is_0_lemma:
"[| m: nat; n: nat |] ==> m #- n = 0 \<longleftrightarrow> m \<le> n"
@@ -538,7 +538,7 @@
apply simp_all
done
-text{*Difference and less-than*}
+text\<open>Difference and less-than\<close>
lemma diff_lt_imp_lt: "[|(k#-i) < (k#-j); i\<in>nat; j\<in>nat; k\<in>nat|] ==> j<i"
apply (erule rev_mp)
--- a/src/ZF/Bin.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/ZF/Bin.thy Thu Jul 23 14:25:05 2015 +0200
@@ -13,7 +13,7 @@
For instance, ~5 div 2 = ~3 and ~5 mod 2 = 1; thus ~5 = (~3)*2 + 1
*)
-section{*Arithmetic on Binary Integers*}
+section\<open>Arithmetic on Binary Integers\<close>
theory Bin
imports Int_ZF Datatype_ZF
@@ -176,8 +176,8 @@
by (induct_tac "v", auto)
-subsubsection{*The Carry and Borrow Functions,
- @{term bin_succ} and @{term bin_pred}*}
+subsubsection\<open>The Carry and Borrow Functions,
+ @{term bin_succ} and @{term bin_pred}\<close>
(*NCons preserves the integer value of its argument*)
lemma integ_of_NCons [simp]:
@@ -199,7 +199,7 @@
done
-subsubsection{*@{term bin_minus}: Unary Negation of Binary Integers*}
+subsubsection\<open>@{term bin_minus}: Unary Negation of Binary Integers\<close>
lemma integ_of_minus: "w \<in> bin ==> integ_of(bin_minus(w)) = $- integ_of(w)"
apply (erule bin.induct)
@@ -207,7 +207,7 @@
done
-subsubsection{*@{term bin_add}: Binary Addition*}
+subsubsection\<open>@{term bin_add}: Binary Addition\<close>
lemma bin_add_Pls [simp]: "w \<in> bin ==> bin_add(Pls,w) = w"
by (unfold bin_add_def, simp)
@@ -255,7 +255,7 @@
done
-subsubsection{*@{term bin_mult}: Binary Multiplication*}
+subsubsection\<open>@{term bin_mult}: Binary Multiplication\<close>
lemma integ_of_mult:
"[| v \<in> bin; w \<in> bin |]
@@ -266,7 +266,7 @@
done
-subsection{*Computations*}
+subsection\<open>Computations\<close>
(** extra rules for bin_succ, bin_pred **)
@@ -351,8 +351,8 @@
done
-subsection{*Simplification Rules for Comparison of Binary Numbers*}
-text{*Thanks to Norbert Voelker*}
+subsection\<open>Simplification Rules for Comparison of Binary Numbers\<close>
+text\<open>Thanks to Norbert Voelker\<close>
(** Equals (=) **)
@@ -695,9 +695,9 @@
ML_file "int_arith.ML"
-subsection {* examples: *}
+subsection \<open>examples:\<close>
-text {* @{text combine_numerals_prod} (products of separate literals) *}
+text \<open>@{text combine_numerals_prod} (products of separate literals)\<close>
lemma "#5 $* x $* #3 = y" apply simp oops
schematic_lemma "y2 $+ ?x42 = y $+ y2" apply simp oops
@@ -741,7 +741,7 @@
lemma "a $+ $-(b$+c) $+ b = d" apply simp oops
lemma "a $+ $-(b$+c) $- b = d" apply simp oops
-text {* negative numerals *}
+text \<open>negative numerals\<close>
lemma "(i $+ j $+ #-2 $+ k) $- (u $+ #5 $+ y) = zz" apply simp oops
lemma "(i $+ j $+ #-3 $+ k) $< u $+ #5 $+ y" apply simp oops
lemma "(i $+ j $+ #3 $+ k) $< u $+ #-6 $+ y" apply simp oops
@@ -749,7 +749,7 @@
lemma "(i $+ j $+ #12 $+ k) $- #-15 = y" apply simp oops
lemma "(i $+ j $+ #-12 $+ k) $- #-15 = y" apply simp oops
-text {* Multiplying separated numerals *}
+text \<open>Multiplying separated numerals\<close>
lemma "#6 $* ($# x $* #2) = uu" apply simp oops
lemma "#4 $* ($# x $* $# x) $* (#2 $* $# x) = uu" apply simp oops
--- a/src/ZF/Bool.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/ZF/Bool.thy Thu Jul 23 14:25:05 2015 +0200
@@ -3,7 +3,7 @@
Copyright 1992 University of Cambridge
*)
-section{*Booleans in Zermelo-Fraenkel Set Theory*}
+section\<open>Booleans in Zermelo-Fraenkel Set Theory\<close>
theory Bool imports pair begin
@@ -15,7 +15,7 @@
two ("2") where
"2 == succ(1)"
-text{*2 is equal to bool, but is used as a number rather than a type.*}
+text\<open>2 is equal to bool, but is used as a number rather than a type.\<close>
definition "bool == {0,1}"
@@ -109,7 +109,7 @@
lemmas bool_typechecks = bool_1I bool_0I cond_type not_type and_type
or_type xor_type
-subsection{*Laws About 'not' *}
+subsection\<open>Laws About 'not'\<close>
lemma not_not [simp]: "a:bool ==> not(not(a)) = a"
by (elim boolE, auto)
@@ -120,7 +120,7 @@
lemma not_or [simp]: "a:bool ==> not(a or b) = not(a) and not(b)"
by (elim boolE, auto)
-subsection{*Laws About 'and' *}
+subsection\<open>Laws About 'and'\<close>
lemma and_absorb [simp]: "a: bool ==> a and a = a"
by (elim boolE, auto)
@@ -135,7 +135,7 @@
(a or b) and c = (a and c) or (b and c)"
by (elim boolE, auto)
-subsection{*Laws About 'or' *}
+subsection\<open>Laws About 'or'\<close>
lemma or_absorb [simp]: "a: bool ==> a or a = a"
by (elim boolE, auto)
--- a/src/ZF/Cardinal.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/ZF/Cardinal.thy Thu Jul 23 14:25:05 2015 +0200
@@ -3,7 +3,7 @@
Copyright 1994 University of Cambridge
*)
-section{*Cardinal Numbers Without the Axiom of Choice*}
+section\<open>Cardinal Numbers Without the Axiom of Choice\<close>
theory Cardinal imports OrderType Finite Nat_ZF Sum begin
@@ -47,8 +47,8 @@
Least (binder "\<mu>" 10)
-subsection{*The Schroeder-Bernstein Theorem*}
-text{*See Davey and Priestly, page 106*}
+subsection\<open>The Schroeder-Bernstein Theorem\<close>
+text\<open>See Davey and Priestly, page 106\<close>
(** Lemma: Banach's Decomposition Theorem **)
@@ -178,7 +178,7 @@
done
-subsection{*lesspoll: contributions by Krzysztof Grabczewski *}
+subsection\<open>lesspoll: contributions by Krzysztof Grabczewski\<close>
lemma lesspoll_not_refl: "~ (i \<prec> i)"
by (simp add: lesspoll_def)
@@ -276,7 +276,7 @@
thus ?thesis using P .
qed
-text{*The proof is almost identical to the one above!*}
+text\<open>The proof is almost identical to the one above!\<close>
lemma Least_le:
assumes P: "P(i)" and i: "Ord(i)" shows "(\<mu> x. P(x)) \<le> i"
proof -
@@ -333,7 +333,7 @@
qed
-subsection{*Basic Properties of Cardinals*}
+subsection\<open>Basic Properties of Cardinals\<close>
(*Not needed for simplification, but helpful below*)
lemma Least_cong: "(!!y. P(y) \<longleftrightarrow> Q(y)) ==> (\<mu> x. P(x)) = (\<mu> x. Q(x))"
@@ -410,7 +410,7 @@
apply (rule Ord_Least)
done
-text{*The cardinals are the initial ordinals.*}
+text\<open>The cardinals are the initial ordinals.\<close>
lemma Card_iff_initial: "Card(K) \<longleftrightarrow> Ord(K) & (\<forall>j. j<K \<longrightarrow> ~ j \<approx> K)"
proof -
{ fix j
@@ -449,10 +449,10 @@
proof (unfold cardinal_def)
show "Card(\<mu> i. i \<approx> A)"
proof (cases "\<exists>i. Ord (i) & i \<approx> A")
- case False thus ?thesis --{*degenerate case*}
+ case False thus ?thesis --\<open>degenerate case\<close>
by (simp add: Least_0 Card_0)
next
- case True --{*real case: @{term A} is isomorphic to some ordinal*}
+ case True --\<open>real case: @{term A} is isomorphic to some ordinal\<close>
then obtain i where i: "Ord(i)" "i \<approx> A" by blast
show ?thesis
proof (rule CardI [OF Ord_Least], rule notI)
@@ -500,7 +500,7 @@
thus ?thesis by simp
qed
-text{*Since we have @{term"|succ(nat)| \<le> |nat|"}, the converse of @{text cardinal_mono} fails!*}
+text\<open>Since we have @{term"|succ(nat)| \<le> |nat|"}, the converse of @{text cardinal_mono} fails!\<close>
lemma cardinal_lt_imp_lt: "[| |i| < |j|; Ord(i); Ord(j) |] ==> i < j"
apply (rule Ord_linear2 [of i j], assumption+)
apply (erule lt_trans2 [THEN lt_irrefl])
@@ -556,7 +556,7 @@
apply (blast intro: Ord_trans)
done
-subsection{*The finite cardinals *}
+subsection\<open>The finite cardinals\<close>
lemma cons_lepoll_consD:
"[| cons(u,A) \<lesssim> cons(v,B); u\<notin>A; v\<notin>B |] ==> A \<lesssim> B"
@@ -591,12 +591,12 @@
case 0 thus ?case by (blast intro!: nat_0_le)
next
case (succ m)
- show ?case using `n \<in> nat`
+ show ?case using \<open>n \<in> nat\<close>
proof (cases rule: natE)
case 0 thus ?thesis using succ
by (simp add: lepoll_def inj_def)
next
- case (succ n') thus ?thesis using succ.hyps ` succ(m) \<lesssim> n`
+ case (succ n') thus ?thesis using succ.hyps \<open> succ(m) \<lesssim> n\<close>
by (blast intro!: succ_leI dest!: succ_lepoll_succD)
qed
qed
@@ -682,7 +682,7 @@
done
-subsection{*The first infinite cardinal: Omega, or nat *}
+subsection\<open>The first infinite cardinal: Omega, or nat\<close>
(*This implies Kunen's Lemma 10.6*)
lemma lt_not_lepoll:
@@ -697,7 +697,7 @@
thus ?thesis by auto
qed
-text{*A slightly weaker version of @{text nat_eqpoll_iff}*}
+text\<open>A slightly weaker version of @{text nat_eqpoll_iff}\<close>
lemma Ord_nat_eqpoll_iff:
assumes i: "Ord(i)" and n: "n \<in> nat" shows "i \<approx> n \<longleftrightarrow> i=n"
using i nat_into_Ord [OF n]
@@ -712,7 +712,7 @@
case gt
hence "~ i \<lesssim> n" using n by (rule lt_not_lepoll)
hence "~ i \<approx> n" using n by (blast intro: eqpoll_imp_lepoll)
- moreover have "i \<noteq> n" using `n<i` by auto
+ moreover have "i \<noteq> n" using \<open>n<i\<close> by auto
ultimately show ?thesis by blast
qed
@@ -740,7 +740,7 @@
by (blast intro: Ord_nat Card_nat ltI lt_Card_imp_lesspoll)
-subsection{*Towards Cardinal Arithmetic *}
+subsection\<open>Towards Cardinal Arithmetic\<close>
(** Congruence laws for successor, cardinal addition and multiplication **)
(*Congruence law for cons under equipollence*)
@@ -817,12 +817,12 @@
done
-subsection{*Lemmas by Krzysztof Grabczewski*}
+subsection\<open>Lemmas by Krzysztof Grabczewski\<close>
(*New proofs using cons_lepoll_cons. Could generalise from succ to cons.*)
-text{*If @{term A} has at most @{term"n+1"} elements and @{term"a \<in> A"}
- then @{term"A-{a}"} has at most @{term n}.*}
+text\<open>If @{term A} has at most @{term"n+1"} elements and @{term"a \<in> A"}
+ then @{term"A-{a}"} has at most @{term n}.\<close>
lemma Diff_sing_lepoll:
"[| a \<in> A; A \<lesssim> succ(n) |] ==> A - {a} \<lesssim> n"
apply (unfold succ_def)
@@ -831,7 +831,7 @@
apply (erule cons_Diff [THEN ssubst], safe)
done
-text{*If @{term A} has at least @{term"n+1"} elements then @{term"A-{a}"} has at least @{term n}.*}
+text\<open>If @{term A} has at least @{term"n+1"} elements then @{term"A-{a}"} has at least @{term n}.\<close>
lemma lepoll_Diff_sing:
assumes A: "succ(n) \<lesssim> A" shows "n \<lesssim> A - {a}"
proof -
@@ -877,7 +877,7 @@
done
-subsection {*Finite and infinite sets*}
+subsection \<open>Finite and infinite sets\<close>
lemma eqpoll_imp_Finite_iff: "A \<approx> B ==> Finite(A) \<longleftrightarrow> Finite(B)"
apply (unfold Finite_def)
@@ -1027,7 +1027,7 @@
lemma Finite_Un_iff [simp]: "Finite(A \<union> B) \<longleftrightarrow> (Finite(A) & Finite(B))"
by (blast intro: subset_Finite Finite_Un)
-text{*The converse must hold too.*}
+text\<open>The converse must hold too.\<close>
lemma Finite_Union: "[| \<forall>y\<in>X. Finite(y); Finite(X) |] ==> Finite(\<Union>(X))"
apply (simp add: Finite_Fin_iff)
apply (rule Fin_UnionI)
@@ -1085,8 +1085,8 @@
apply (blast intro: elim: equalityCE)
done
-text{*I don't know why, but if the premise is expressed using meta-connectives
-then the simplifier cannot prove it automatically in conditional rewriting.*}
+text\<open>I don't know why, but if the premise is expressed using meta-connectives
+then the simplifier cannot prove it automatically in conditional rewriting.\<close>
lemma Finite_RepFun_iff:
"(\<forall>x y. f(x)=f(y) \<longrightarrow> x=y) ==> Finite(RepFun(A,f)) \<longleftrightarrow> Finite(A)"
by (blast intro: Finite_RepFun Finite_RepFun_iff_lemma [of _ f])
@@ -1119,7 +1119,7 @@
next
case (succ x)
hence wfx: "\<And>Z. Z = 0 \<or> (\<exists>z\<in>Z. \<forall>y. z \<in> y \<and> z \<in> x \<and> y \<in> x \<and> z \<in> x \<longrightarrow> y \<notin> Z)"
- by (simp add: wf_on_def wf_def) --{*not easy to erase the duplicate @{term"z \<in> x"}!*}
+ by (simp add: wf_on_def wf_def) --\<open>not easy to erase the duplicate @{term"z \<in> x"}!\<close>
show ?case
proof (rule wf_onI)
fix Z u
--- a/src/ZF/CardinalArith.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/ZF/CardinalArith.thy Thu Jul 23 14:25:05 2015 +0200
@@ -3,7 +3,7 @@
Copyright 1994 University of Cambridge
*)
-section{*Cardinal Arithmetic Without the Axiom of Choice*}
+section\<open>Cardinal Arithmetic Without the Axiom of Choice\<close>
theory CardinalArith imports Cardinal OrderArith ArithSimp Finite begin
@@ -28,15 +28,15 @@
definition
jump_cardinal :: "i=>i" where
- --{*This def is more complex than Kunen's but it more easily proved to
- be a cardinal*}
+ --\<open>This def is more complex than Kunen's but it more easily proved to
+ be a cardinal\<close>
"jump_cardinal(K) ==
\<Union>X\<in>Pow(K). {z. r \<in> Pow(K*K), well_ord(X,r) & z = ordertype(X,r)}"
definition
csucc :: "i=>i" where
- --{*needed because @{term "jump_cardinal(K)"} might not be the successor
- of @{term K}*}
+ --\<open>needed because @{term "jump_cardinal(K)"} might not be the successor
+ of @{term K}\<close>
"csucc(K) == LEAST L. Card(L) & K<L"
notation (xsymbols)
@@ -87,14 +87,14 @@
done
-subsection{*Cardinal addition*}
+subsection\<open>Cardinal addition\<close>
-text{*Note: Could omit proving the algebraic laws for cardinal addition and
+text\<open>Note: Could omit proving the algebraic laws for cardinal addition and
multiplication. On finite cardinals these operations coincide with
addition and multiplication of natural numbers; on infinite cardinals they
-coincide with union (maximum). Either way we get most laws for free.*}
+coincide with union (maximum). Either way we get most laws for free.\<close>
-subsubsection{*Cardinal addition is commutative*}
+subsubsection\<open>Cardinal addition is commutative\<close>
lemma sum_commute_eqpoll: "A+B \<approx> B+A"
proof (unfold eqpoll_def, rule exI)
@@ -107,7 +107,7 @@
apply (rule sum_commute_eqpoll [THEN cardinal_cong])
done
-subsubsection{*Cardinal addition is associative*}
+subsubsection\<open>Cardinal addition is associative\<close>
lemma sum_assoc_eqpoll: "(A+B)+C \<approx> A+(B+C)"
apply (unfold eqpoll_def)
@@ -115,7 +115,7 @@
apply (rule sum_assoc_bij)
done
-text{*Unconditional version requires AC*}
+text\<open>Unconditional version requires AC\<close>
lemma well_ord_cadd_assoc:
assumes i: "well_ord(i,ri)" and j: "well_ord(j,rj)" and k: "well_ord(k,rk)"
shows "(i \<oplus> j) \<oplus> k = i \<oplus> (j \<oplus> k)"
@@ -130,7 +130,7 @@
qed
-subsubsection{*0 is the identity for addition*}
+subsubsection\<open>0 is the identity for addition\<close>
lemma sum_0_eqpoll: "0+A \<approx> A"
apply (unfold eqpoll_def)
@@ -143,7 +143,7 @@
apply (simp add: sum_0_eqpoll [THEN cardinal_cong] Card_cardinal_eq)
done
-subsubsection{*Addition by another cardinal*}
+subsubsection\<open>Addition by another cardinal\<close>
lemma sum_lepoll_self: "A \<lesssim> A+B"
proof (unfold lepoll_def, rule exI)
@@ -165,7 +165,7 @@
by (blast intro: le_trans)
qed
-subsubsection{*Monotonicity of addition*}
+subsubsection\<open>Monotonicity of addition\<close>
lemma sum_lepoll_mono:
"[| A \<lesssim> C; B \<lesssim> D |] ==> A + B \<lesssim> C + D"
@@ -186,7 +186,7 @@
apply (blast intro: sum_lepoll_mono subset_imp_lepoll)
done
-subsubsection{*Addition of finite cardinals is "ordinary" addition*}
+subsubsection\<open>Addition of finite cardinals is "ordinary" addition\<close>
lemma sum_succ_eqpoll: "succ(A)+B \<approx> succ(A+B)"
apply (unfold eqpoll_def)
@@ -222,9 +222,9 @@
qed
-subsection{*Cardinal multiplication*}
+subsection\<open>Cardinal multiplication\<close>
-subsubsection{*Cardinal multiplication is commutative*}
+subsubsection\<open>Cardinal multiplication is commutative\<close>
lemma prod_commute_eqpoll: "A*B \<approx> B*A"
apply (unfold eqpoll_def)
@@ -238,7 +238,7 @@
apply (rule prod_commute_eqpoll [THEN cardinal_cong])
done
-subsubsection{*Cardinal multiplication is associative*}
+subsubsection\<open>Cardinal multiplication is associative\<close>
lemma prod_assoc_eqpoll: "(A*B)*C \<approx> A*(B*C)"
apply (unfold eqpoll_def)
@@ -246,7 +246,7 @@
apply (rule prod_assoc_bij)
done
-text{*Unconditional version requires AC*}
+text\<open>Unconditional version requires AC\<close>
lemma well_ord_cmult_assoc:
assumes i: "well_ord(i,ri)" and j: "well_ord(j,rj)" and k: "well_ord(k,rk)"
shows "(i \<otimes> j) \<otimes> k = i \<otimes> (j \<otimes> k)"
@@ -260,7 +260,7 @@
finally show "|i * j| * k \<approx> i * |j * k|" .
qed
-subsubsection{*Cardinal multiplication distributes over addition*}
+subsubsection\<open>Cardinal multiplication distributes over addition\<close>
lemma sum_prod_distrib_eqpoll: "(A+B)*C \<approx> (A*C)+(B*C)"
apply (unfold eqpoll_def)
@@ -281,7 +281,7 @@
finally show "|i + j| * k \<approx> |i * k| + |j * k|" .
qed
-subsubsection{*Multiplication by 0 yields 0*}
+subsubsection\<open>Multiplication by 0 yields 0\<close>
lemma prod_0_eqpoll: "0*A \<approx> 0"
apply (unfold eqpoll_def)
@@ -292,7 +292,7 @@
lemma cmult_0 [simp]: "0 \<otimes> i = 0"
by (simp add: cmult_def prod_0_eqpoll [THEN cardinal_cong])
-subsubsection{*1 is the identity for multiplication*}
+subsubsection\<open>1 is the identity for multiplication\<close>
lemma prod_singleton_eqpoll: "{x}*A \<approx> A"
apply (unfold eqpoll_def)
@@ -305,7 +305,7 @@
apply (simp add: prod_singleton_eqpoll [THEN cardinal_cong] Card_cardinal_eq)
done
-subsection{*Some inequalities for multiplication*}
+subsection\<open>Some inequalities for multiplication\<close>
lemma prod_square_lepoll: "A \<lesssim> A*A"
apply (unfold lepoll_def inj_def)
@@ -322,7 +322,7 @@
apply (blast intro: well_ord_rmult well_ord_Memrel Card_is_Ord)
done
-subsubsection{*Multiplication by a non-zero cardinal*}
+subsubsection\<open>Multiplication by a non-zero cardinal\<close>
lemma prod_lepoll_self: "b \<in> B ==> A \<lesssim> A*B"
apply (unfold lepoll_def inj_def)
@@ -339,7 +339,7 @@
apply (blast intro: prod_lepoll_self ltD)
done
-subsubsection{*Monotonicity of multiplication*}
+subsubsection\<open>Monotonicity of multiplication\<close>
lemma prod_lepoll_mono:
"[| A \<lesssim> C; B \<lesssim> D |] ==> A * B \<lesssim> C * D"
@@ -360,7 +360,7 @@
apply (blast intro: prod_lepoll_mono subset_imp_lepoll)
done
-subsection{*Multiplication of finite cardinals is "ordinary" multiplication*}
+subsection\<open>Multiplication of finite cardinals is "ordinary" multiplication\<close>
lemma prod_succ_eqpoll: "succ(A)*B \<approx> B + A*B"
apply (unfold eqpoll_def)
@@ -403,7 +403,7 @@
by (blast intro: sum_lepoll_mono sum_lepoll_prod lepoll_trans lepoll_refl)
-subsection{*Infinite Cardinals are Limit Ordinals*}
+subsection\<open>Infinite Cardinals are Limit Ordinals\<close>
(*This proof is modelled upon one assuming nat<=A, with injection
\<lambda>z\<in>cons(u,A). if z=u then 0 else if z \<in> nat then succ(z) else z
@@ -484,7 +484,7 @@
apply (rule pred_subset)
done
-subsubsection{*Establishing the well-ordering*}
+subsubsection\<open>Establishing the well-ordering\<close>
lemma well_ord_csquare:
assumes K: "Ord(K)" shows "well_ord(K*K, csquare_rel(K))"
@@ -496,7 +496,7 @@
using K by (blast intro: well_ord_rmult well_ord_Memrel)
qed
-subsubsection{*Characterising initial segments of the well-ordering*}
+subsubsection\<open>Characterising initial segments of the well-ordering\<close>
lemma csquareD:
"[| <<x,y>, <z,z>> \<in> csquare_rel(K); x<K; y<K; z<K |] ==> x \<le> z & y \<le> z"
@@ -537,7 +537,7 @@
subset_Un_iff2 [THEN iff_sym] OrdmemD)
done
-subsubsection{*The cardinality of initial segments*}
+subsubsection\<open>The cardinality of initial segments\<close>
lemma ordermap_z_lt:
"[| Limit(K); x<K; y<K; z=succ(x \<union> y) |] ==>
@@ -551,7 +551,7 @@
apply (blast intro!: Un_upper1_le Un_upper2_le Ord_ordermap elim!: ltE)+
done
-text{*Kunen: "each @{term"\<langle>x,y\<rangle> \<in> K \<times> K"} has no more than @{term"z \<times> z"} predecessors..." (page 29) *}
+text\<open>Kunen: "each @{term"\<langle>x,y\<rangle> \<in> K \<times> K"} has no more than @{term"z \<times> z"} predecessors..." (page 29)\<close>
lemma ordermap_csquare_le:
assumes K: "Limit(K)" and x: "x<K" and y: " y<K"
defines "z \<equiv> succ(x \<union> y)"
@@ -582,7 +582,7 @@
finally show "ordermap(K \<times> K, csquare_rel(K)) ` \<langle>x,y\<rangle> \<lesssim> |succ(z)| \<times> |succ(z)|" .
qed
-text{*Kunen: "... so the order type is @{text"\<le>"} K" *}
+text\<open>Kunen: "... so the order type is @{text"\<le>"} K"\<close>
lemma ordertype_csquare_le:
assumes IK: "InfCard(K)" and eq: "\<And>y. y\<in>K \<Longrightarrow> InfCard(y) \<Longrightarrow> y \<otimes> y = y"
shows "ordertype(K*K, csquare_rel(K)) \<le> K"
@@ -685,7 +685,7 @@
lemma Inf_Card_is_InfCard: "[| Card(i); ~ Finite(i) |] ==> InfCard(i)"
by (simp add: InfCard_def Card_is_Ord [THEN nat_le_infinite_Ord])
-subsubsection{*Toward's Kunen's Corollary 10.13 (1)*}
+subsubsection\<open>Toward's Kunen's Corollary 10.13 (1)\<close>
lemma InfCard_le_cmult_eq: "[| InfCard(K); L \<le> K; 0<L |] ==> K \<otimes> L = K"
apply (rule le_anti_sym)
@@ -734,9 +734,9 @@
might be InfCard(K) ==> |list(K)| = K.
*)
-subsection{*For Every Cardinal Number There Exists A Greater One*}
+subsection\<open>For Every Cardinal Number There Exists A Greater One\<close>
-text{*This result is Kunen's Theorem 10.16, which would be trivial using AC*}
+text\<open>This result is Kunen's Theorem 10.16, which would be trivial using AC\<close>
lemma Ord_jump_cardinal: "Ord(jump_cardinal(K))"
apply (unfold jump_cardinal_def)
@@ -793,7 +793,7 @@
apply (blast intro: Card_jump_cardinal_lemma [THEN mem_irrefl])
done
-subsection{*Basic Properties of Successor Cardinals*}
+subsection\<open>Basic Properties of Successor Cardinals\<close>
lemma csucc_basic: "Ord(K) ==> Card(csucc(K)) & K < csucc(K)"
apply (unfold csucc_def)
@@ -834,7 +834,7 @@
lt_csucc [THEN leI, THEN [2] le_trans])
-subsubsection{*Removing elements from a finite set decreases its cardinality*}
+subsubsection\<open>Removing elements from a finite set decreases its cardinality\<close>
lemma Finite_imp_cardinal_cons [simp]:
assumes FA: "Finite(A)" and a: "a\<notin>A" shows "|cons(a,A)| = succ(|A|)"
@@ -918,7 +918,7 @@
qed
-subsubsection{*Theorems by Krzysztof Grabczewski, proofs by lcp*}
+subsubsection\<open>Theorems by Krzysztof Grabczewski, proofs by lcp\<close>
lemmas nat_implies_well_ord = nat_into_Ord [THEN well_ord_Memrel]
--- a/src/ZF/Cardinal_AC.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/ZF/Cardinal_AC.thy Thu Jul 23 14:25:05 2015 +0200
@@ -5,18 +5,18 @@
These results help justify infinite-branching datatypes
*)
-section{*Cardinal Arithmetic Using AC*}
+section\<open>Cardinal Arithmetic Using AC\<close>
theory Cardinal_AC imports CardinalArith Zorn begin
-subsection{*Strengthened Forms of Existing Theorems on Cardinals*}
+subsection\<open>Strengthened Forms of Existing Theorems on Cardinals\<close>
lemma cardinal_eqpoll: "|A| \<approx> A"
apply (rule AC_well_ord [THEN exE])
apply (erule well_ord_cardinal_eqpoll)
done
-text{*The theorem @{term "||A|| = |A|"} *}
+text\<open>The theorem @{term "||A|| = |A|"}\<close>
lemmas cardinal_idem = cardinal_eqpoll [THEN cardinal_cong, simp]
lemma cardinal_eqE: "|X| = |Y| ==> X \<approx> Y"
@@ -65,7 +65,7 @@
done
-subsection {*The relationship between cardinality and le-pollence*}
+subsection \<open>The relationship between cardinality and le-pollence\<close>
lemma Card_le_imp_lepoll:
assumes "|A| \<le> |B|" shows "A \<lesssim> B"
@@ -113,7 +113,7 @@
by (blast intro: lt_Ord Card_le_imp_lepoll Ord_cardinal_le le_trans)
-subsection{*Other Applications of AC*}
+subsection\<open>Other Applications of AC\<close>
lemma surj_implies_inj:
assumes f: "f \<in> surj(X,Y)" shows "\<exists>g. g \<in> inj(Y,X)"
@@ -129,7 +129,7 @@
qed
qed
-text{*Kunen's Lemma 10.20*}
+text\<open>Kunen's Lemma 10.20\<close>
lemma surj_implies_cardinal_le:
assumes f: "f \<in> surj(X,Y)" shows "|Y| \<le> |X|"
proof (rule lepoll_imp_Card_le)
@@ -138,7 +138,7 @@
by (auto simp add: lepoll_def)
qed
-text{*Kunen's Lemma 10.21*}
+text\<open>Kunen's Lemma 10.21\<close>
lemma cardinal_UN_le:
assumes K: "InfCard(K)"
shows "(!!i. i\<in>K ==> |X(i)| \<le> K) ==> |\<Union>i\<in>K. X(i)| \<le> K"
@@ -171,14 +171,14 @@
finally show "(\<Union>i\<in>K. X(i)) \<lesssim> K" .
qed
-text{*The same again, using @{term csucc}*}
+text\<open>The same again, using @{term csucc}\<close>
lemma cardinal_UN_lt_csucc:
"[| InfCard(K); \<And>i. i\<in>K \<Longrightarrow> |X(i)| < csucc(K) |]
==> |\<Union>i\<in>K. X(i)| < csucc(K)"
by (simp add: Card_lt_csucc_iff cardinal_UN_le InfCard_is_Card Card_cardinal)
-text{*The same again, for a union of ordinals. In use, j(i) is a bit like rank(i),
- the least ordinal j such that i:Vfrom(A,j). *}
+text\<open>The same again, for a union of ordinals. In use, j(i) is a bit like rank(i),
+ the least ordinal j such that i:Vfrom(A,j).\<close>
lemma cardinal_UN_Ord_lt_csucc:
"[| InfCard(K); \<And>i. i\<in>K \<Longrightarrow> j(i) < csucc(K) |]
==> (\<Union>i\<in>K. j(i)) < csucc(K)"
@@ -189,11 +189,11 @@
done
-subsection{*The Main Result for Infinite-Branching Datatypes*}
+subsection\<open>The Main Result for Infinite-Branching Datatypes\<close>
-text{*As above, but the index set need not be a cardinal. Work
+text\<open>As above, but the index set need not be a cardinal. Work
backwards along the injection from @{term W} into @{term K}, given
-that @{term"W\<noteq>0"}.*}
+that @{term"W\<noteq>0"}.\<close>
lemma inj_UN_subset:
assumes f: "f \<in> inj(A,B)" and a: "a \<in> A"
@@ -222,7 +222,7 @@
note lt_subset_trans [OF _ _ OU, trans]
show ?thesis
proof (cases "W=0")
- case True --{*solve the easy 0 case*}
+ case True --\<open>solve the easy 0 case\<close>
thus ?thesis by (simp add: CK Card_is_Ord Card_csucc Ord_0_lt_csucc)
next
case False
--- a/src/ZF/Coind/Language.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/ZF/Coind/Language.thy Thu Jul 23 14:25:05 2015 +0200
@@ -6,7 +6,7 @@
theory Language imports Main begin
-text{*these really can't be definitions without losing the abstraction*}
+text\<open>these really can't be definitions without losing the abstraction\<close>
axiomatization
Const :: i and (* Abstract type of constants *)
--- a/src/ZF/Constructible/AC_in_L.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/ZF/Constructible/AC_in_L.thy Thu Jul 23 14:25:05 2015 +0200
@@ -2,13 +2,13 @@
Author: Lawrence C Paulson, Cambridge University Computer Laboratory
*)
-section {* The Axiom of Choice Holds in L! *}
+section \<open>The Axiom of Choice Holds in L!\<close>
theory AC_in_L imports Formula Separation begin
-subsection{*Extending a Wellordering over a List -- Lexicographic Power*}
+subsection\<open>Extending a Wellordering over a List -- Lexicographic Power\<close>
-text{*This could be moved into a library.*}
+text\<open>This could be moved into a library.\<close>
consts
rlist :: "[i,i]=>i"
@@ -31,13 +31,13 @@
type_intros list.intros
-subsubsection{*Type checking*}
+subsubsection\<open>Type checking\<close>
lemmas rlist_type = rlist.dom_subset
lemmas field_rlist = rlist_type [THEN field_rel_subset]
-subsubsection{*Linearity*}
+subsubsection\<open>Linearity\<close>
lemma rlist_Nil_Cons [intro]:
"[|a \<in> A; l \<in> list(A)|] ==> <[], Cons(a,l)> \<in> rlist(A, r)"
@@ -64,7 +64,7 @@
done
}
note yConsCase = this
- show ?case using `ys \<in> list(A)`
+ show ?case using \<open>ys \<in> list(A)\<close>
by (cases rule: list.cases) (simp_all add: Cons rlist_Nil_Cons yConsCase)
qed
}
@@ -72,9 +72,9 @@
qed
-subsubsection{*Well-foundedness*}
+subsubsection\<open>Well-foundedness\<close>
-text{*Nothing preceeds Nil in this ordering.*}
+text\<open>Nothing preceeds Nil in this ordering.\<close>
inductive_cases rlist_NilE: " <l,[]> \<in> rlist(A,r)"
inductive_cases rlist_ConsE: " <l', Cons(x,l)> \<in> rlist(A,r)"
@@ -139,9 +139,9 @@
done
-subsection{*An Injection from Formulas into the Natural Numbers*}
+subsection\<open>An Injection from Formulas into the Natural Numbers\<close>
-text{*There is a well-known bijection between @{term "nat*nat"} and @{term
+text\<open>There is a well-known bijection between @{term "nat*nat"} and @{term
nat} given by the expression f(m,n) = triangle(m+n) + m, where triangle(k)
enumerates the triangular numbers and can be defined by triangle(0)=0,
triangle(succ(k)) = succ(k + triangle(k)). Some small amount of effort is
@@ -151,10 +151,10 @@
However, this result merely states that there is a bijection between the two
sets. It provides no means of naming a specific bijection. Therefore, we
conduct the proofs under the assumption that a bijection exists. The simplest
-way to organize this is to use a locale.*}
+way to organize this is to use a locale.\<close>
-text{*Locale for any arbitrary injection between @{term "nat*nat"}
- and @{term nat}*}
+text\<open>Locale for any arbitrary injection between @{term "nat*nat"}
+ and @{term nat}\<close>
locale Nat_Times_Nat =
fixes fn
assumes fn_inj: "fn \<in> inj(nat*nat, nat)"
@@ -214,7 +214,7 @@
InfCard_nat [THEN InfCard_square_eqpoll, THEN eqpoll_imp_lepoll]
-text{*Not needed--but interesting?*}
+text\<open>Not needed--but interesting?\<close>
theorem formula_lepoll_nat: "formula \<lesssim> nat"
apply (insert nat_times_nat_lepoll_nat)
apply (unfold lepoll_def)
@@ -222,31 +222,31 @@
done
-subsection{*Defining the Wellordering on @{term "DPow(A)"}*}
+subsection\<open>Defining the Wellordering on @{term "DPow(A)"}\<close>
-text{*The objective is to build a wellordering on @{term "DPow(A)"} from a
+text\<open>The objective is to build a wellordering on @{term "DPow(A)"} from a
given one on @{term A}. We first introduce wellorderings for environments,
which are lists built over @{term "A"}. We combine it with the enumeration of
formulas. The order type of the resulting wellordering gives us a map from
(environment, formula) pairs into the ordinals. For each member of @{term
-"DPow(A)"}, we take the minimum such ordinal.*}
+"DPow(A)"}, we take the minimum such ordinal.\<close>
definition
env_form_r :: "[i,i,i]=>i" where
- --{*wellordering on (environment, formula) pairs*}
+ --\<open>wellordering on (environment, formula) pairs\<close>
"env_form_r(f,r,A) ==
rmult(list(A), rlist(A, r),
formula, measure(formula, enum(f)))"
definition
env_form_map :: "[i,i,i,i]=>i" where
- --{*map from (environment, formula) pairs to ordinals*}
+ --\<open>map from (environment, formula) pairs to ordinals\<close>
"env_form_map(f,r,A,z)
== ordermap(list(A) * formula, env_form_r(f,r,A)) ` z"
definition
DPow_ord :: "[i,i,i,i,i]=>o" where
- --{*predicate that holds if @{term k} is a valid index for @{term X}*}
+ --\<open>predicate that holds if @{term k} is a valid index for @{term X}\<close>
"DPow_ord(f,r,A,X,k) ==
\<exists>env \<in> list(A). \<exists>p \<in> formula.
arity(p) \<le> succ(length(env)) &
@@ -255,12 +255,12 @@
definition
DPow_least :: "[i,i,i,i]=>i" where
- --{*function yielding the smallest index for @{term X}*}
+ --\<open>function yielding the smallest index for @{term X}\<close>
"DPow_least(f,r,A,X) == \<mu> k. DPow_ord(f,r,A,X,k)"
definition
DPow_r :: "[i,i,i]=>i" where
- --{*a wellordering on @{term "DPow(A)"}*}
+ --\<open>a wellordering on @{term "DPow(A)"}\<close>
"DPow_r(f,r,A) == measure(DPow(A), DPow_least(f,r,A))"
@@ -324,16 +324,16 @@
by (simp add: DPow_r_def measure_def, blast)
-subsection{*Limit Construction for Well-Orderings*}
+subsection\<open>Limit Construction for Well-Orderings\<close>
-text{*Now we work towards the transfinite definition of wellorderings for
+text\<open>Now we work towards the transfinite definition of wellorderings for
@{term "Lset(i)"}. We assume as an inductive hypothesis that there is a family
-of wellorderings for smaller ordinals.*}
+of wellorderings for smaller ordinals.\<close>
definition
rlimit :: "[i,i=>i]=>i" where
- --{*Expresses the wellordering at limit ordinals. The conditional
- lets us remove the premise @{term "Limit(i)"} from some theorems.*}
+ --\<open>Expresses the wellordering at limit ordinals. The conditional
+ lets us remove the premise @{term "Limit(i)"} from some theorems.\<close>
"rlimit(i,r) ==
if Limit(i) then
{z: Lset(i) * Lset(i).
@@ -344,8 +344,8 @@
definition
Lset_new :: "i=>i" where
- --{*This constant denotes the set of elements introduced at level
- @{term "succ(i)"}*}
+ --\<open>This constant denotes the set of elements introduced at level
+ @{term "succ(i)"}\<close>
"Lset_new(i) == {x \<in> Lset(succ(i)). lrank(x) = i}"
lemma Limit_Lset_eq2:
@@ -412,7 +412,7 @@
done
-subsection{*Transfinite Definition of the Wellordering on @{term "L"}*}
+subsection\<open>Transfinite Definition of the Wellordering on @{term "L"}\<close>
definition
L_r :: "[i, i] => i" where
@@ -420,15 +420,15 @@
transrec3(i, 0, \<lambda>x r. DPow_r(f, r, Lset(x)),
\<lambda>x r. rlimit(x, \<lambda>y. r`y))"
-subsubsection{*The Corresponding Recursion Equations*}
+subsubsection\<open>The Corresponding Recursion Equations\<close>
lemma [simp]: "L_r(f,0) = 0"
by (simp add: L_r_def)
lemma [simp]: "L_r(f, succ(i)) = DPow_r(f, L_r(f,i), Lset(i))"
by (simp add: L_r_def)
-text{*The limit case is non-trivial because of the distinction between
-object-level and meta-level abstraction.*}
+text\<open>The limit case is non-trivial because of the distinction between
+object-level and meta-level abstraction.\<close>
lemma [simp]: "Limit(i) ==> L_r(f,i) = rlimit(i, L_r(f))"
by (simp cong: rlimit_cong add: transrec3_Limit L_r_def ltD)
@@ -454,8 +454,8 @@
done
-text{*Every constructible set is well-ordered! Therefore the Wellordering Theorem and
- the Axiom of Choice hold in @{term L}!!*}
+text\<open>Every constructible set is well-ordered! Therefore the Wellordering Theorem and
+ the Axiom of Choice hold in @{term L}!!\<close>
theorem L_implies_AC: assumes x: "L(x)" shows "\<exists>r. well_ord(x,r)"
using Transset_Lset x
apply (simp add: Transset_def L_def)
@@ -474,8 +474,8 @@
by (blast intro: well_ord_imp_relativized)
qed
-text{*In order to prove @{term" \<exists>r[L]. wellordered(L,x,r)"}, it's necessary to know
+text\<open>In order to prove @{term" \<exists>r[L]. wellordered(L,x,r)"}, it's necessary to know
that @{term r} is actually constructible. It follows from the assumption ``@{term V} equals @{term L''},
-but this reasoning doesn't appear to work in Isabelle.*}
+but this reasoning doesn't appear to work in Isabelle.\<close>
end
--- a/src/ZF/Constructible/DPow_absolute.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/ZF/Constructible/DPow_absolute.thy Thu Jul 23 14:25:05 2015 +0200
@@ -2,18 +2,18 @@
Author: Lawrence C Paulson, Cambridge University Computer Laboratory
*)
-section {*Absoluteness for the Definable Powerset Function*}
+section \<open>Absoluteness for the Definable Powerset Function\<close>
theory DPow_absolute imports Satisfies_absolute begin
-subsection{*Preliminary Internalizations*}
+subsection\<open>Preliminary Internalizations\<close>
-subsubsection{*The Operator @{term is_formula_rec}*}
+subsubsection\<open>The Operator @{term is_formula_rec}\<close>
-text{*The three arguments of @{term p} are always 2, 1, 0. It is buried
- within 11 quantifiers!!*}
+text\<open>The three arguments of @{term p} are always 2, 1, 0. It is buried
+ within 11 quantifiers!!\<close>
(* is_formula_rec :: "[i=>o, [i,i,i]=>o, i, i] => o"
"is_formula_rec(M,MH,p,z) ==
@@ -77,7 +77,7 @@
done
-subsubsection{*The Operator @{term is_satisfies}*}
+subsubsection\<open>The Operator @{term is_satisfies}\<close>
(* is_satisfies(M,A,p,z) == is_formula_rec (M, satisfies_MH(M,A), p, z) *)
definition
@@ -109,7 +109,7 @@
done
-subsection {*Relativization of the Operator @{term DPow'}*}
+subsection \<open>Relativization of the Operator @{term DPow'}\<close>
lemma DPow'_eq:
"DPow'(A) = {z . ep \<in> list(A) * formula,
@@ -118,8 +118,8 @@
by (simp add: DPow'_def, blast)
-text{*Relativize the use of @{term sats} within @{term DPow'}
-(the comprehension).*}
+text\<open>Relativize the use of @{term sats} within @{term DPow'}
+(the comprehension).\<close>
definition
is_DPow_sats :: "[i=>o,i,i,i,i] => o" where
"is_DPow_sats(M,A,env,p,x) ==
@@ -142,7 +142,7 @@
by (simp add: DPow_sats_abs transM [of _ A])
-subsubsection{*The Operator @{term is_DPow_sats}, Internalized*}
+subsubsection\<open>The Operator @{term is_DPow_sats}, Internalized\<close>
(* is_DPow_sats(M,A,env,p,x) ==
\<forall>n1[M]. \<forall>e[M]. \<forall>sp[M].
@@ -184,7 +184,7 @@
done
-subsection{*A Locale for Relativizing the Operator @{term DPow'}*}
+subsection\<open>A Locale for Relativizing the Operator @{term DPow'}\<close>
locale M_DPow = M_satisfies +
assumes sep:
@@ -219,7 +219,7 @@
apply (fast intro: rep' sep' univalent_pair_eq)
done
-text{*Relativization of the Operator @{term DPow'}*}
+text\<open>Relativization of the Operator @{term DPow'}\<close>
definition
is_DPow' :: "[i=>o,i,i] => o" where
"is_DPow'(M,A,Z) ==
@@ -238,9 +238,9 @@
done
-subsection{*Instantiating the Locale @{text M_DPow}*}
+subsection\<open>Instantiating the Locale @{text M_DPow}\<close>
-subsubsection{*The Instance of Separation*}
+subsubsection\<open>The Instance of Separation\<close>
lemma DPow_separation:
"[| L(A); env \<in> list(A); p \<in> formula |]
@@ -253,7 +253,7 @@
-subsubsection{*The Instance of Replacement*}
+subsubsection\<open>The Instance of Replacement\<close>
lemma DPow_replacement_Reflects:
"REFLECTS [\<lambda>x. \<exists>u[L]. u \<in> B &
@@ -287,7 +287,7 @@
done
-subsubsection{*Actually Instantiating the Locale*}
+subsubsection\<open>Actually Instantiating the Locale\<close>
lemma M_DPow_axioms_L: "M_DPow_axioms(L)"
apply (rule M_DPow_axioms.intro)
@@ -304,10 +304,10 @@
and DPow'_abs [intro, simp] = M_DPow.DPow'_abs [OF M_DPow_L]
-subsubsection{*The Operator @{term is_Collect}*}
+subsubsection\<open>The Operator @{term is_Collect}\<close>
-text{*The formula @{term is_P} has one free variable, 0, and it is
-enclosed within a single quantifier.*}
+text\<open>The formula @{term is_P} has one free variable, 0, and it is
+enclosed within a single quantifier.\<close>
(* is_Collect :: "[i=>o,i,i=>o,i] => o"
"is_Collect(M,A,P,z) == \<forall>x[M]. x \<in> z \<longleftrightarrow> x \<in> A & P(x)" *)
@@ -342,8 +342,8 @@
by (simp add: sats_Collect_fm [OF is_P_iff_sats])
-text{*The second argument of @{term is_P} gives it direct access to @{term x},
- which is essential for handling free variable references.*}
+text\<open>The second argument of @{term is_P} gives it direct access to @{term x},
+ which is essential for handling free variable references.\<close>
theorem Collect_reflection:
assumes is_P_reflection:
"!!h f g. REFLECTS[\<lambda>x. is_P(L, f(x), g(x)),
@@ -355,10 +355,10 @@
done
-subsubsection{*The Operator @{term is_Replace}*}
+subsubsection\<open>The Operator @{term is_Replace}\<close>
-text{*BEWARE! The formula @{term is_P} has free variables 0, 1
- and not the usual 1, 0! It is enclosed within two quantifiers.*}
+text\<open>BEWARE! The formula @{term is_P} has free variables 0, 1
+ and not the usual 1, 0! It is enclosed within two quantifiers.\<close>
(* is_Replace :: "[i=>o,i,[i,i]=>o,i] => o"
"is_Replace(M,A,P,z) == \<forall>u[M]. u \<in> z \<longleftrightarrow> (\<exists>x[M]. x\<in>A & P(x,u))" *)
@@ -395,8 +395,8 @@
by (simp add: sats_Replace_fm [OF is_P_iff_sats])
-text{*The second argument of @{term is_P} gives it direct access to @{term x},
- which is essential for handling free variable references.*}
+text\<open>The second argument of @{term is_P} gives it direct access to @{term x},
+ which is essential for handling free variable references.\<close>
theorem Replace_reflection:
assumes is_P_reflection:
"!!h f g. REFLECTS[\<lambda>x. is_P(L, f(x), g(x), h(x)),
@@ -409,7 +409,7 @@
-subsubsection{*The Operator @{term is_DPow'}, Internalized*}
+subsubsection\<open>The Operator @{term is_DPow'}, Internalized\<close>
(* "is_DPow'(M,A,Z) ==
\<forall>X[M]. X \<in> Z \<longleftrightarrow>
@@ -454,7 +454,7 @@
done
-subsection{*A Locale for Relativizing the Operator @{term Lset}*}
+subsection\<open>A Locale for Relativizing the Operator @{term Lset}\<close>
definition
transrec_body :: "[i=>o,i,i,i,i] => o" where
@@ -506,13 +506,13 @@
done
-text{*Relativization of the Operator @{term Lset}*}
+text\<open>Relativization of the Operator @{term Lset}\<close>
definition
is_Lset :: "[i=>o, i, i] => o" where
- --{*We can use the term language below because @{term is_Lset} will
+ --\<open>We can use the term language below because @{term is_Lset} will
not have to be internalized: it isn't used in any instance of
- separation.*}
+ separation.\<close>
"is_Lset(M,a,z) == is_transrec(M, %x f u. u = (\<Union>y\<in>x. DPow'(f`y)), a, z)"
lemma (in M_Lset) Lset_abs:
@@ -531,9 +531,9 @@
done
-subsection{*Instantiating the Locale @{text M_Lset}*}
+subsection\<open>Instantiating the Locale @{text M_Lset}\<close>
-subsubsection{*The First Instance of Replacement*}
+subsubsection\<open>The First Instance of Replacement\<close>
lemma strong_rep_Reflects:
"REFLECTS [\<lambda>u. \<exists>v[L]. v \<in> B & (\<exists>gy[L].
@@ -553,7 +553,7 @@
done
-subsubsection{*The Second Instance of Replacement*}
+subsubsection\<open>The Second Instance of Replacement\<close>
lemma transrec_rep_Reflects:
"REFLECTS [\<lambda>x. \<exists>v[L]. v \<in> B &
@@ -570,8 +570,8 @@
is_DPow'(##Lset(i),gy,z), r) &
big_union(##Lset(i),r,u), mr, v, y))]"
apply (simp only: rex_setclass_is_bex [symmetric])
- --{*Convert @{text "\<exists>y\<in>Lset(i)"} to @{text "\<exists>y[##Lset(i)]"} within the body
- of the @{term is_wfrec} application. *}
+ --\<open>Convert @{text "\<exists>y\<in>Lset(i)"} to @{text "\<exists>y[##Lset(i)]"} within the body
+ of the @{term is_wfrec} application.\<close>
apply (intro FOL_reflections function_reflections
is_wfrec_reflection Replace_reflection DPow'_reflection)
done
@@ -593,7 +593,7 @@
done
-subsubsection{*Actually Instantiating @{text M_Lset}*}
+subsubsection\<open>Actually Instantiating @{text M_Lset}\<close>
lemma M_Lset_axioms_L: "M_Lset_axioms(L)"
apply (rule M_Lset_axioms.intro)
@@ -606,12 +606,12 @@
apply (rule M_Lset_axioms_L)
done
-text{*Finally: the point of the whole theory!*}
+text\<open>Finally: the point of the whole theory!\<close>
lemmas Lset_closed = M_Lset.Lset_closed [OF M_Lset_L]
and Lset_abs = M_Lset.Lset_abs [OF M_Lset_L]
-subsection{*The Notion of Constructible Set*}
+subsection\<open>The Notion of Constructible Set\<close>
definition
constructible :: "[i=>o,i] => o" where
--- a/src/ZF/Constructible/Datatype_absolute.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/ZF/Constructible/Datatype_absolute.thy Thu Jul 23 14:25:05 2015 +0200
@@ -2,12 +2,12 @@
Author: Lawrence C Paulson, Cambridge University Computer Laboratory
*)
-section {*Absoluteness Properties for Recursive Datatypes*}
+section \<open>Absoluteness Properties for Recursive Datatypes\<close>
theory Datatype_absolute imports Formula WF_absolute begin
-subsection{*The lfp of a continuous function can be expressed as a union*}
+subsection\<open>The lfp of a continuous function can be expressed as a union\<close>
definition
directed :: "i=>o" where
@@ -78,7 +78,7 @@
intro: lfp_subset_Union Union_subset_lfp)
-subsubsection{*Some Standard Datatype Constructions Preserve Continuity*}
+subsubsection\<open>Some Standard Datatype Constructions Preserve Continuity\<close>
lemma contin_imp_mono: "[|X\<subseteq>Y; contin(F)|] ==> F(X) \<subseteq> F(Y)"
apply (simp add: contin_def)
@@ -111,7 +111,7 @@
-subsection {*Absoluteness for "Iterates"*}
+subsection \<open>Absoluteness for "Iterates"\<close>
definition
iterates_MH :: "[i=>o, [i,i]=>o, i, i, i, i] => o" where
@@ -169,7 +169,7 @@
done
-subsection {*lists without univ*}
+subsection \<open>lists without univ\<close>
lemmas datatype_univs = Inl_in_univ Inr_in_univ
Pair_in_univ nat_into_univ A_into_univ
@@ -184,7 +184,7 @@
lemma list_fun_contin: "contin(\<lambda>X. {0} + A*X)"
by (intro sum_contin prod_contin id_contin const_contin)
-text{*Re-expresses lists using sum and product*}
+text\<open>Re-expresses lists using sum and product\<close>
lemma list_eq_lfp2: "list(A) = lfp(univ(A), \<lambda>X. {0} + A*X)"
apply (simp add: list_def)
apply (rule equalityI)
@@ -193,7 +193,7 @@
apply (clarify, subst lfp_unfold [OF list_fun_bnd_mono])
apply (simp add: Nil_def Cons_def)
apply blast
-txt{*Opposite inclusion*}
+txt\<open>Opposite inclusion\<close>
apply (rule lfp_lowerbound)
prefer 2 apply (rule lfp_subset)
apply (clarify, subst lfp_unfold [OF list.bnd_mono])
@@ -202,7 +202,7 @@
dest: lfp_subset [THEN subsetD])
done
-text{*Re-expresses lists using "iterates", no univ.*}
+text\<open>Re-expresses lists using "iterates", no univ.\<close>
lemma list_eq_Union:
"list(A) = (\<Union>n\<in>nat. (\<lambda>X. {0} + A*X) ^ n (0))"
by (simp add: list_eq_lfp2 lfp_eq_Union list_fun_bnd_mono list_fun_contin)
@@ -219,7 +219,7 @@
by (simp add: is_list_functor_def singleton_0 nat_into_M)
-subsection {*formulas without univ*}
+subsection \<open>formulas without univ\<close>
lemma formula_fun_bnd_mono:
"bnd_mono(univ(0), \<lambda>X. ((nat*nat) + (nat*nat)) + (X*X + X))"
@@ -234,7 +234,7 @@
by (intro sum_contin prod_contin id_contin const_contin)
-text{*Re-expresses formulas using sum and product*}
+text\<open>Re-expresses formulas using sum and product\<close>
lemma formula_eq_lfp2:
"formula = lfp(univ(0), \<lambda>X. ((nat*nat) + (nat*nat)) + (X*X + X))"
apply (simp add: formula_def)
@@ -244,7 +244,7 @@
apply (clarify, subst lfp_unfold [OF formula_fun_bnd_mono])
apply (simp add: Member_def Equal_def Nand_def Forall_def)
apply blast
-txt{*Opposite inclusion*}
+txt\<open>Opposite inclusion\<close>
apply (rule lfp_lowerbound)
prefer 2 apply (rule lfp_subset, clarify)
apply (subst lfp_unfold [OF formula.bnd_mono, simplified])
@@ -253,7 +253,7 @@
apply (blast intro: datatype_univs dest: lfp_subset [THEN subsetD])+
done
-text{*Re-expresses formulas using "iterates", no univ.*}
+text\<open>Re-expresses formulas using "iterates", no univ.\<close>
lemma formula_eq_Union:
"formula =
(\<Union>n\<in>nat. (\<lambda>X. ((nat*nat) + (nat*nat)) + (X*X + X)) ^ n (0))"
@@ -277,7 +277,7 @@
by (simp add: is_formula_functor_def)
-subsection{*@{term M} Contains the List and Formula Datatypes*}
+subsection\<open>@{term M} Contains the List and Formula Datatypes\<close>
definition
list_N :: "[i,i] => i" where
@@ -290,8 +290,8 @@
"Cons(a,l) \<in> list_N(A,succ(n)) \<longleftrightarrow> a\<in>A & l \<in> list_N(A,n)"
by (simp add: list_N_def Cons_def)
-text{*These two aren't simprules because they reveal the underlying
-list representation.*}
+text\<open>These two aren't simprules because they reveal the underlying
+list representation.\<close>
lemma list_N_0: "list_N(A,0) = 0"
by (simp add: list_N_def)
@@ -325,8 +325,8 @@
apply (blast intro: list_imp_list_N)
done
-text{*Express @{term list_rec} without using @{term rank} or @{term Vset},
-neither of which is absolute.*}
+text\<open>Express @{term list_rec} without using @{term rank} or @{term Vset},
+neither of which is absolute.\<close>
lemma (in M_trivial) list_rec_eq:
"l \<in> list(A) ==>
list_rec(a,g,l) =
@@ -356,7 +356,7 @@
is_list :: "[i=>o,i,i] => o" where
"is_list(M,A,Z) == \<forall>l[M]. l \<in> Z \<longleftrightarrow> mem_list(M,A,l)"
-subsubsection{*Towards Absoluteness of @{term formula_rec}*}
+subsubsection\<open>Towards Absoluteness of @{term formula_rec}\<close>
consts depth :: "i=>i"
primrec
@@ -389,8 +389,8 @@
"Forall(x) \<in> formula_N(succ(n)) \<longleftrightarrow> x \<in> formula_N(n)"
by (simp add: formula_N_def Forall_def)
-text{*These two aren't simprules because they reveal the underlying
-formula representation.*}
+text\<open>These two aren't simprules because they reveal the underlying
+formula representation.\<close>
lemma formula_N_0: "formula_N(0) = 0"
by (simp add: formula_N_def)
@@ -430,7 +430,7 @@
done
-text{*This result and the next are unused.*}
+text\<open>This result and the next are unused.\<close>
lemma formula_N_mono [rule_format]:
"[| m \<in> nat; n \<in> nat |] ==> m\<le>n \<longrightarrow> formula_N(m) \<subseteq> formula_N(n)"
apply (rule_tac m = m and n = n in diff_induct)
@@ -476,7 +476,7 @@
"M(l) ==> iterates_replacement(M, %l t. is_tl(M,l,t), l)"
-subsubsection{*Absoluteness of the List Construction*}
+subsubsection\<open>Absoluteness of the List Construction\<close>
lemma (in M_datatypes) list_replacement2':
"M(A) ==> strong_replacement(M, \<lambda>n y. n\<in>nat & y = (\<lambda>X. {0} + A * X)^n (0))"
@@ -493,7 +493,7 @@
list_replacement2' relation1_def
iterates_closed [of "is_list_functor(M,A)"])
-text{*WARNING: use only with @{text "dest:"} or with variables fixed!*}
+text\<open>WARNING: use only with @{text "dest:"} or with variables fixed!\<close>
lemmas (in M_datatypes) list_into_M = transM [OF _ list_closed]
lemma (in M_datatypes) list_N_abs [simp]:
@@ -524,7 +524,7 @@
apply (rule M_equalityI, simp_all)
done
-subsubsection{*Absoluteness of Formulas*}
+subsubsection\<open>Absoluteness of Formulas\<close>
lemma (in M_datatypes) formula_replacement2':
"strong_replacement(M, \<lambda>n y. n\<in>nat & y = (\<lambda>X. ((nat*nat) + (nat*nat)) + (X*X + X))^n (0))"
@@ -574,9 +574,9 @@
done
-subsection{*Absoluteness for @{text \<epsilon>}-Closure: the @{term eclose} Operator*}
+subsection\<open>Absoluteness for @{text \<epsilon>}-Closure: the @{term eclose} Operator\<close>
-text{*Re-expresses eclose using "iterates"*}
+text\<open>Re-expresses eclose using "iterates"\<close>
lemma eclose_eq_Union:
"eclose(A) = (\<Union>n\<in>nat. Union^n (A))"
apply (simp add: eclose_def)
@@ -645,9 +645,9 @@
done
-subsection {*Absoluteness for @{term transrec}*}
+subsection \<open>Absoluteness for @{term transrec}\<close>
-text{* @{prop "transrec(a,H) \<equiv> wfrec(Memrel(eclose({a})), a, H)"} *}
+text\<open>@{prop "transrec(a,H) \<equiv> wfrec(Memrel(eclose({a})), a, H)"}\<close>
definition
is_transrec :: "[i=>o, [i,i,i]=>o, i, i] => o" where
@@ -663,9 +663,9 @@
upair(M,a,a,sa) & is_eclose(M,sa,esa) & membership(M,esa,mesa) &
wfrec_replacement(M,MH,mesa)"
-text{*The condition @{term "Ord(i)"} lets us use the simpler
+text\<open>The condition @{term "Ord(i)"} lets us use the simpler
@{text "trans_wfrec_abs"} rather than @{text "trans_wfrec_abs"},
- which I haven't even proved yet. *}
+ which I haven't even proved yet.\<close>
theorem (in M_eclose) transrec_abs:
"[|transrec_replacement(M,MH,i); relation2(M,MH,H);
Ord(i); M(i); M(z);
@@ -684,7 +684,7 @@
transrec_def eclose_sing_Ord_eq wf_Memrel trans_Memrel relation_Memrel)
-text{*Helps to prove instances of @{term transrec_replacement}*}
+text\<open>Helps to prove instances of @{term transrec_replacement}\<close>
lemma (in M_eclose) transrec_replacementI:
"[|M(a);
strong_replacement (M,
@@ -694,8 +694,8 @@
by (simp add: transrec_replacement_def wfrec_replacement_def)
-subsection{*Absoluteness for the List Operator @{term length}*}
-text{*But it is never used.*}
+subsection\<open>Absoluteness for the List Operator @{term length}\<close>
+text\<open>But it is never used.\<close>
definition
is_length :: "[i=>o,i,i,i] => o" where
@@ -714,13 +714,13 @@
dest: list_N_imp_length_lt)
done
-text{*Proof is trivial since @{term length} returns natural numbers.*}
+text\<open>Proof is trivial since @{term length} returns natural numbers.\<close>
lemma (in M_trivial) length_closed [intro,simp]:
"l \<in> list(A) ==> M(length(l))"
by (simp add: nat_into_M)
-subsection {*Absoluteness for the List Operator @{term nth}*}
+subsection \<open>Absoluteness for the List Operator @{term nth}\<close>
lemma nth_eq_hd_iterates_tl [rule_format]:
"xs \<in> list(A) ==> \<forall>n \<in> nat. nth(n,xs) = hd' (tl'^n (xs))"
@@ -737,7 +737,7 @@
apply (simp add: tl'_Cons tl'_closed)
done
-text{*Immediate by type-checking*}
+text\<open>Immediate by type-checking\<close>
lemma (in M_datatypes) nth_closed [intro,simp]:
"[|xs \<in> list(A); n \<in> nat; M(A)|] ==> M(nth(n,xs))"
apply (case_tac "n < length(xs)")
@@ -761,11 +761,11 @@
done
-subsection{*Relativization and Absoluteness for the @{term formula} Constructors*}
+subsection\<open>Relativization and Absoluteness for the @{term formula} Constructors\<close>
definition
is_Member :: "[i=>o,i,i,i] => o" where
- --{* because @{term "Member(x,y) \<equiv> Inl(Inl(\<langle>x,y\<rangle>))"}*}
+ --\<open>because @{term "Member(x,y) \<equiv> Inl(Inl(\<langle>x,y\<rangle>))"}\<close>
"is_Member(M,x,y,Z) ==
\<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) & is_Inl(M,p,u) & is_Inl(M,u,Z)"
@@ -779,7 +779,7 @@
definition
is_Equal :: "[i=>o,i,i,i] => o" where
- --{* because @{term "Equal(x,y) \<equiv> Inl(Inr(\<langle>x,y\<rangle>))"}*}
+ --\<open>because @{term "Equal(x,y) \<equiv> Inl(Inr(\<langle>x,y\<rangle>))"}\<close>
"is_Equal(M,x,y,Z) ==
\<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) & is_Inr(M,p,u) & is_Inl(M,u,Z)"
@@ -792,7 +792,7 @@
definition
is_Nand :: "[i=>o,i,i,i] => o" where
- --{* because @{term "Nand(x,y) \<equiv> Inr(Inl(\<langle>x,y\<rangle>))"}*}
+ --\<open>because @{term "Nand(x,y) \<equiv> Inr(Inl(\<langle>x,y\<rangle>))"}\<close>
"is_Nand(M,x,y,Z) ==
\<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) & is_Inl(M,p,u) & is_Inr(M,u,Z)"
@@ -805,7 +805,7 @@
definition
is_Forall :: "[i=>o,i,i] => o" where
- --{* because @{term "Forall(x) \<equiv> Inr(Inr(p))"}*}
+ --\<open>because @{term "Forall(x) \<equiv> Inr(Inr(p))"}\<close>
"is_Forall(M,p,Z) == \<exists>u[M]. is_Inr(M,p,u) & is_Inr(M,u,Z)"
lemma (in M_trivial) Forall_abs [simp]:
@@ -817,20 +817,20 @@
-subsection {*Absoluteness for @{term formula_rec}*}
+subsection \<open>Absoluteness for @{term formula_rec}\<close>
definition
formula_rec_case :: "[[i,i]=>i, [i,i]=>i, [i,i,i,i]=>i, [i,i]=>i, i, i] => i" where
- --{* the instance of @{term formula_case} in @{term formula_rec}*}
+ --\<open>the instance of @{term formula_case} in @{term formula_rec}\<close>
"formula_rec_case(a,b,c,d,h) ==
formula_case (a, b,
\<lambda>u v. c(u, v, h ` succ(depth(u)) ` u,
h ` succ(depth(v)) ` v),
\<lambda>u. d(u, h ` succ(depth(u)) ` u))"
-text{*Unfold @{term formula_rec} to @{term formula_rec_case}.
+text\<open>Unfold @{term formula_rec} to @{term formula_rec_case}.
Express @{term formula_rec} without using @{term rank} or @{term Vset},
-neither of which is absolute.*}
+neither of which is absolute.\<close>
lemma (in M_trivial) formula_rec_eq:
"p \<in> formula ==>
formula_rec(a,b,c,d,p) =
@@ -838,20 +838,20 @@
\<lambda>x h. Lambda (formula, formula_rec_case(a,b,c,d,h))) ` p"
apply (simp add: formula_rec_case_def)
apply (induct_tac p)
- txt{*Base case for @{term Member}*}
+ txt\<open>Base case for @{term Member}\<close>
apply (subst transrec, simp add: formula.intros)
- txt{*Base case for @{term Equal}*}
+ txt\<open>Base case for @{term Equal}\<close>
apply (subst transrec, simp add: formula.intros)
- txt{*Inductive step for @{term Nand}*}
+ txt\<open>Inductive step for @{term Nand}\<close>
apply (subst transrec)
apply (simp add: succ_Un_distrib formula.intros)
-txt{*Inductive step for @{term Forall}*}
+txt\<open>Inductive step for @{term Forall}\<close>
apply (subst transrec)
apply (simp add: formula_imp_formula_N formula.intros)
done
-subsubsection{*Absoluteness for the Formula Operator @{term depth}*}
+subsubsection\<open>Absoluteness for the Formula Operator @{term depth}\<close>
definition
is_depth :: "[i=>o,i,i] => o" where
@@ -870,18 +870,18 @@
dest: formula_N_imp_depth_lt)
done
-text{*Proof is trivial since @{term depth} returns natural numbers.*}
+text\<open>Proof is trivial since @{term depth} returns natural numbers.\<close>
lemma (in M_trivial) depth_closed [intro,simp]:
"p \<in> formula ==> M(depth(p))"
by (simp add: nat_into_M)
-subsubsection{*@{term is_formula_case}: relativization of @{term formula_case}*}
+subsubsection\<open>@{term is_formula_case}: relativization of @{term formula_case}\<close>
definition
is_formula_case ::
"[i=>o, [i,i,i]=>o, [i,i,i]=>o, [i,i,i]=>o, [i,i]=>o, i, i] => o" where
- --{*no constraint on non-formulas*}
+ --\<open>no constraint on non-formulas\<close>
"is_formula_case(M, is_a, is_b, is_c, is_d, p, z) ==
(\<forall>x[M]. \<forall>y[M]. finite_ordinal(M,x) \<longrightarrow> finite_ordinal(M,y) \<longrightarrow>
is_Member(M,x,y,p) \<longrightarrow> is_a(x,y,z)) &
@@ -911,18 +911,18 @@
by (erule formula.cases, simp_all)
-subsubsection {*Absoluteness for @{term formula_rec}: Final Results*}
+subsubsection \<open>Absoluteness for @{term formula_rec}: Final Results\<close>
definition
is_formula_rec :: "[i=>o, [i,i,i]=>o, i, i] => o" where
- --{* predicate to relativize the functional @{term formula_rec}*}
+ --\<open>predicate to relativize the functional @{term formula_rec}\<close>
"is_formula_rec(M,MH,p,z) ==
\<exists>dp[M]. \<exists>i[M]. \<exists>f[M]. finite_ordinal(M,dp) & is_depth(M,p,dp) &
successor(M,dp,i) & fun_apply(M,f,p,z) & is_transrec(M,MH,i,f)"
-text{*Sufficient conditions to relativize the instance of @{term formula_case}
- in @{term formula_rec}*}
+text\<open>Sufficient conditions to relativize the instance of @{term formula_case}
+ in @{term formula_rec}\<close>
lemma (in M_datatypes) Relation1_formula_rec_case:
"[|Relation2(M, nat, nat, is_a, a);
Relation2(M, nat, nat, is_b, b);
@@ -939,9 +939,9 @@
done
-text{*This locale packages the premises of the following theorems,
+text\<open>This locale packages the premises of the following theorems,
which is the normal purpose of locales. It doesn't accumulate
- constraints on the class @{term M}, as in most of this deveopment.*}
+ constraints on the class @{term M}, as in most of this deveopment.\<close>
locale Formula_Rec = M_eclose +
fixes a and is_a and b and is_b and c and is_c and d and is_d and MH
defines
@@ -995,7 +995,7 @@
by (simp add: transrec_closed [OF fr_replace MH_rel2]
nat_into_M formula_rec_lam_closed)
-text{*The main two results: @{term formula_rec} is absolute for @{term M}.*}
+text\<open>The main two results: @{term formula_rec} is absolute for @{term M}.\<close>
theorem (in Formula_Rec) formula_rec_closed:
"p \<in> formula ==> M(formula_rec(a,b,c,d,p))"
by (simp add: formula_rec_eq fr_transrec_closed
--- a/src/ZF/Constructible/Formula.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/ZF/Constructible/Formula.thy Thu Jul 23 14:25:05 2015 +0200
@@ -2,14 +2,14 @@
Author: Lawrence C Paulson, Cambridge University Computer Laboratory
*)
-section {* First-Order Formulas and the Definition of the Class L *}
+section \<open>First-Order Formulas and the Definition of the Class L\<close>
theory Formula imports Main begin
-subsection{*Internalized formulas of FOL*}
+subsection\<open>Internalized formulas of FOL\<close>
-text{*De Bruijn representation.
- Unbound variables get their denotations from an environment.*}
+text\<open>De Bruijn representation.
+ Unbound variables get their denotations from an environment.\<close>
consts formula :: i
datatype
@@ -109,7 +109,7 @@
declare satisfies.simps [simp del]
-subsection{*Dividing line between primitive and derived connectives*}
+subsection\<open>Dividing line between primitive and derived connectives\<close>
lemma sats_Neg_iff [simp]:
"env \<in> list(A)
@@ -142,7 +142,7 @@
by (simp add: Exists_def)
-subsubsection{*Derived rules to help build up formulas*}
+subsubsection\<open>Derived rules to help build up formulas\<close>
lemma mem_iff_sats:
"[| nth(i,env) = x; nth(j,env) = y; env \<in> list(A)|]
@@ -195,7 +195,7 @@
bex_iff_sats
-subsection{*Arity of a Formula: Maximum Free de Bruijn Index*}
+subsection\<open>Arity of a Formula: Maximum Free de Bruijn Index\<close>
consts arity :: "i=>i"
primrec
@@ -249,7 +249,7 @@
done
-subsection{*Renaming Some de Bruijn Variables*}
+subsection\<open>Renaming Some de Bruijn Variables\<close>
definition
incr_var :: "[i,i]=>i" where
@@ -284,9 +284,9 @@
lemma incr_bv_type [TC]: "p \<in> formula ==> incr_bv(p) \<in> nat -> formula"
by (induct_tac p, simp_all)
-text{*Obviously, @{term DPow} is closed under complements and finite
+text\<open>Obviously, @{term DPow} is closed under complements and finite
intersections and unions. Needs an inductive lemma to allow two lists of
-parameters to be combined.*}
+parameters to be combined.\<close>
lemma sats_incr_bv_iff [rule_format]:
"[| p \<in> formula; env \<in> list(A); x \<in> A |]
@@ -325,11 +325,11 @@
succ_Un_distrib [symmetric] incr_var_lt incr_var_le
Un_commute incr_var_lemma Arith.pred_def nat_imp_quasinat
split: split_nat_case)
- txt{*the Forall case reduces to linear arithmetic*}
+ txt\<open>the Forall case reduces to linear arithmetic\<close>
prefer 2
apply clarify
apply (blast dest: lt_trans1)
-txt{*left with the And case*}
+txt\<open>left with the And case\<close>
apply safe
apply (blast intro: incr_And_lemma lt_trans1)
apply (subst incr_And_lemma)
@@ -338,7 +338,7 @@
done
-subsection{*Renaming all but the First de Bruijn Variable*}
+subsection\<open>Renaming all but the First de Bruijn Variable\<close>
definition
incr_bv1 :: "i => i" where
@@ -389,9 +389,9 @@
-subsection{*Definable Powerset*}
+subsection\<open>Definable Powerset\<close>
-text{*The definable powerset operation: Kunen's definition VI 1.1, page 165.*}
+text\<open>The definable powerset operation: Kunen's definition VI 1.1, page 165.\<close>
definition
DPow :: "i => i" where
"DPow(A) == {X \<in> Pow(A).
@@ -404,7 +404,7 @@
==> {x\<in>A. sats(A, p, Cons(x,env))} \<in> DPow(A)"
by (simp add: DPow_def, blast)
-text{*With this rule we can specify @{term p} later.*}
+text\<open>With this rule we can specify @{term p} later.\<close>
lemma DPowI2 [rule_format]:
"[|\<forall>x\<in>A. P(x) \<longleftrightarrow> sats(A, p, Cons(x,env));
env \<in> list(A); p \<in> formula; arity(p) \<le> succ(length(env))|]
@@ -482,10 +482,10 @@
apply (blast intro: cons_in_DPow)
done
-text{*@{term DPow} is not monotonic. For example, let @{term A} be some
+text\<open>@{term DPow} is not monotonic. For example, let @{term A} be some
non-constructible set of natural numbers, and let @{term B} be @{term nat}.
Then @{term "A<=B"} and obviously @{term "A \<in> DPow(A)"} but @{term "A \<notin>
-DPow(B)"}.*}
+DPow(B)"}.\<close>
(*This may be true but the proof looks difficult, requiring relativization
lemma DPow_insert: "DPow (cons(a,A)) = DPow(A) \<union> {cons(a,X) . X \<in> DPow(A)}"
@@ -503,17 +503,17 @@
done
-subsection{*Internalized Formulas for the Ordinals*}
+subsection\<open>Internalized Formulas for the Ordinals\<close>
-text{*The @{text sats} theorems below differ from the usual form in that they
+text\<open>The @{text sats} theorems below differ from the usual form in that they
include an element of absoluteness. That is, they relate internalized
formulas to real concepts such as the subset relation, rather than to the
relativized concepts defined in theory @{text Relative}. This lets us prove
the theorem as @{text Ords_in_DPow} without first having to instantiate the
locale @{text M_trivial}. Note that the present theory does not even take
-@{text Relative} as a parent.*}
+@{text Relative} as a parent.\<close>
-subsubsection{*The subset relation*}
+subsubsection\<open>The subset relation\<close>
definition
subset_fm :: "[i,i]=>i" where
@@ -534,7 +534,7 @@
apply (blast intro: nth_type)
done
-subsubsection{*Transitive sets*}
+subsubsection\<open>Transitive sets\<close>
definition
transset_fm :: "i=>i" where
@@ -555,7 +555,7 @@
apply (blast intro: nth_type)
done
-subsubsection{*Ordinals*}
+subsubsection\<open>Ordinals\<close>
definition
ordinal_fm :: "i=>i" where
@@ -577,8 +577,8 @@
apply (blast intro: nth_type)
done
-text{*The subset consisting of the ordinals is definable. Essential lemma for
-@{text Ord_in_Lset}. This result is the objective of the present subsection.*}
+text\<open>The subset consisting of the ordinals is definable. Essential lemma for
+@{text Ord_in_Lset}. This result is the objective of the present subsection.\<close>
theorem Ords_in_DPow: "Transset(A) ==> {x \<in> A. Ord(x)} \<in> DPow(A)"
apply (simp add: DPow_def Collect_subset)
apply (rule_tac x=Nil in bexI)
@@ -587,17 +587,17 @@
done
-subsection{* Constant Lset: Levels of the Constructible Universe *}
+subsection\<open>Constant Lset: Levels of the Constructible Universe\<close>
definition
Lset :: "i=>i" where
"Lset(i) == transrec(i, %x f. \<Union>y\<in>x. DPow(f`y))"
definition
- L :: "i=>o" where --{*Kunen's definition VI 1.5, page 167*}
+ L :: "i=>o" where --\<open>Kunen's definition VI 1.5, page 167\<close>
"L(x) == \<exists>i. Ord(i) & x \<in> Lset(i)"
-text{*NOT SUITABLE FOR REWRITING -- RECURSIVE!*}
+text\<open>NOT SUITABLE FOR REWRITING -- RECURSIVE!\<close>
lemma Lset: "Lset(i) = (\<Union>j\<in>i. DPow(Lset(j)))"
by (subst Lset_def [THEN def_transrec], simp)
@@ -609,7 +609,7 @@
apply (blast intro: elim: equalityE)
done
-subsubsection{* Transitivity *}
+subsubsection\<open>Transitivity\<close>
lemma elem_subset_in_DPow: "[|X \<in> A; X \<subseteq> A|] ==> X \<in> DPow(A)"
apply (simp add: Transset_def DPow_def)
@@ -629,7 +629,7 @@
apply (blast intro: elem_subset_in_DPow dest: DPowD)
done
-text{*Kunen's VI 1.6 (a)*}
+text\<open>Kunen's VI 1.6 (a)\<close>
lemma Transset_Lset: "Transset(Lset(i))"
apply (rule_tac a=i in eps_induct)
apply (subst Lset)
@@ -641,9 +641,9 @@
apply (simp add: Transset_def)
done
-subsubsection{* Monotonicity *}
+subsubsection\<open>Monotonicity\<close>
-text{*Kunen's VI 1.6 (b)*}
+text\<open>Kunen's VI 1.6 (b)\<close>
lemma Lset_mono [rule_format]:
"\<forall>j. i<=j \<longrightarrow> Lset(i) \<subseteq> Lset(j)"
proof (induct i rule: eps_induct, intro allI impI)
@@ -654,7 +654,7 @@
by (force simp add: Lset [of x] Lset [of j])
qed
-text{*This version lets us remove the premise @{term "Ord(i)"} sometimes.*}
+text\<open>This version lets us remove the premise @{term "Ord(i)"} sometimes.\<close>
lemma Lset_mono_mem [rule_format]:
"\<forall>j. i \<in> j \<longrightarrow> Lset(i) \<subseteq> Lset(j)"
proof (induct i rule: eps_induct, intro allI impI)
@@ -667,11 +667,11 @@
qed
-text{*Useful with Reflection to bump up the ordinal*}
+text\<open>Useful with Reflection to bump up the ordinal\<close>
lemma subset_Lset_ltD: "[|A \<subseteq> Lset(i); i < j|] ==> A \<subseteq> Lset(j)"
by (blast dest: ltD [THEN Lset_mono_mem])
-subsubsection{* 0, successor and limit equations for Lset *}
+subsubsection\<open>0, successor and limit equations for Lset\<close>
lemma Lset_0 [simp]: "Lset(0) = 0"
by (subst Lset, blast)
@@ -696,17 +696,17 @@
lemma Lset_Union [simp]: "Lset(\<Union>(X)) = (\<Union>y\<in>X. Lset(y))"
apply (subst Lset)
apply (rule equalityI)
- txt{*first inclusion*}
+ txt\<open>first inclusion\<close>
apply (rule UN_least)
apply (erule UnionE)
apply (rule subset_trans)
apply (erule_tac [2] UN_upper, subst Lset, erule UN_upper)
-txt{*opposite inclusion*}
+txt\<open>opposite inclusion\<close>
apply (rule UN_least)
apply (subst Lset, blast)
done
-subsubsection{* Lset applied to Limit ordinals *}
+subsubsection\<open>Lset applied to Limit ordinals\<close>
lemma Limit_Lset_eq:
"Limit(i) ==> Lset(i) = (\<Union>y\<in>i. Lset(y))"
@@ -726,7 +726,7 @@
apply (blast intro: ltI Limit_is_Ord)
done
-subsubsection{* Basic closure properties *}
+subsubsection\<open>Basic closure properties\<close>
lemma zero_in_Lset: "y \<in> x ==> 0 \<in> Lset(x)"
by (subst Lset, blast intro: empty_in_DPow)
@@ -738,22 +738,22 @@
done
-subsection{*Constructible Ordinals: Kunen's VI 1.9 (b)*}
+subsection\<open>Constructible Ordinals: Kunen's VI 1.9 (b)\<close>
lemma Ords_of_Lset_eq: "Ord(i) ==> {x\<in>Lset(i). Ord(x)} = i"
apply (erule trans_induct3)
apply (simp_all add: Lset_succ Limit_Lset_eq Limit_Union_eq)
-txt{*The successor case remains.*}
+txt\<open>The successor case remains.\<close>
apply (rule equalityI)
-txt{*First inclusion*}
+txt\<open>First inclusion\<close>
apply clarify
apply (erule Ord_linear_lt, assumption)
apply (blast dest: DPow_imp_subset ltD notE [OF notin_Lset])
apply blast
apply (blast dest: ltD)
-txt{*Opposite inclusion, @{term "succ(x) \<subseteq> DPow(Lset(x)) \<inter> ON"}*}
+txt\<open>Opposite inclusion, @{term "succ(x) \<subseteq> DPow(Lset(x)) \<inter> ON"}\<close>
apply auto
-txt{*Key case: *}
+txt\<open>Key case:\<close>
apply (erule subst, rule Ords_in_DPow [OF Transset_Lset])
apply (blast intro: elem_subset_in_DPow dest: OrdmemD elim: equalityE)
apply (blast intro: Ord_in_Ord)
@@ -772,7 +772,7 @@
lemma Ord_in_L: "Ord(i) ==> L(i)"
by (simp add: L_def, blast intro: Ord_in_Lset)
-subsubsection{* Unions *}
+subsubsection\<open>Unions\<close>
lemma Union_in_Lset:
"X \<in> Lset(i) ==> \<Union>(X) \<in> Lset(succ(i))"
@@ -780,7 +780,7 @@
apply (rule LsetI [OF succI1])
apply (simp add: Transset_def DPow_def)
apply (intro conjI, blast)
-txt{*Now to create the formula @{term "\<exists>y. y \<in> X \<and> x \<in> y"} *}
+txt\<open>Now to create the formula @{term "\<exists>y. y \<in> X \<and> x \<in> y"}\<close>
apply (rule_tac x="Cons(X,Nil)" in bexI)
apply (rule_tac x="Exists(And(Member(0,2), Member(1,0)))" in bexI)
apply typecheck
@@ -790,7 +790,7 @@
theorem Union_in_L: "L(X) ==> L(\<Union>(X))"
by (simp add: L_def, blast dest: Union_in_Lset)
-subsubsection{* Finite sets and ordered pairs *}
+subsubsection\<open>Finite sets and ordered pairs\<close>
lemma singleton_in_Lset: "a \<in> Lset(i) ==> {a} \<in> Lset(succ(i))"
by (simp add: Lset_succ singleton_in_DPow)
@@ -808,7 +808,7 @@
lemmas Lset_UnI1 = Un_upper1 [THEN Lset_mono [THEN subsetD]]
lemmas Lset_UnI2 = Un_upper2 [THEN Lset_mono [THEN subsetD]]
-text{*Hard work is finding a single @{term"j \<in> i"} such that @{term"{a,b} \<subseteq> Lset(j)"}*}
+text\<open>Hard work is finding a single @{term"j \<in> i"} such that @{term"{a,b} \<subseteq> Lset(j)"}\<close>
lemma doubleton_in_LLimit:
"[| a \<in> Lset(i); b \<in> Lset(i); Limit(i) |] ==> {a,b} \<in> Lset(i)"
apply (erule Limit_LsetE, assumption)
@@ -825,19 +825,19 @@
lemma Pair_in_LLimit:
"[| a \<in> Lset(i); b \<in> Lset(i); Limit(i) |] ==> <a,b> \<in> Lset(i)"
-txt{*Infer that a, b occur at ordinals x,xa < i.*}
+txt\<open>Infer that a, b occur at ordinals x,xa < i.\<close>
apply (erule Limit_LsetE, assumption)
apply (erule Limit_LsetE, assumption)
-txt{*Infer that @{term"succ(succ(x \<union> xa)) < i"} *}
+txt\<open>Infer that @{term"succ(succ(x \<union> xa)) < i"}\<close>
apply (blast intro: lt_Ord lt_LsetI [OF Pair_in_Lset]
Lset_UnI1 Lset_UnI2 Limit_has_succ Un_least_lt)
done
-text{*The rank function for the constructible universe*}
+text\<open>The rank function for the constructible universe\<close>
definition
- lrank :: "i=>i" where --{*Kunen's definition VI 1.7*}
+ lrank :: "i=>i" where --\<open>Kunen's definition VI 1.7\<close>
"lrank(x) == \<mu> i. x \<in> Lset(succ(i))"
lemma L_I: "[|x \<in> Lset(i); Ord(i)|] ==> L(x)"
@@ -858,9 +858,9 @@
apply (blast intro: ltI Limit_is_Ord lt_trans)
done
-text{*Kunen's VI 1.8. The proof is much harder than the text would
+text\<open>Kunen's VI 1.8. The proof is much harder than the text would
suggest. For a start, it needs the previous lemma, which is proved by
-induction.*}
+induction.\<close>
lemma Lset_iff_lrank_lt: "Ord(i) ==> x \<in> Lset(i) \<longleftrightarrow> L(x) & lrank(x) < i"
apply (simp add: L_def, auto)
apply (blast intro: Lset_lrank_lt)
@@ -873,7 +873,7 @@
lemma Lset_succ_lrank_iff [simp]: "x \<in> Lset(succ(lrank(x))) \<longleftrightarrow> L(x)"
by (simp add: Lset_iff_lrank_lt)
-text{*Kunen's VI 1.9 (a)*}
+text\<open>Kunen's VI 1.9 (a)\<close>
lemma lrank_of_Ord: "Ord(i) ==> lrank(i) = i"
apply (unfold lrank_def)
apply (rule Least_equality)
@@ -884,10 +884,10 @@
done
-text{*This is lrank(lrank(a)) = lrank(a) *}
+text\<open>This is lrank(lrank(a)) = lrank(a)\<close>
declare Ord_lrank [THEN lrank_of_Ord, simp]
-text{*Kunen's VI 1.10 *}
+text\<open>Kunen's VI 1.10\<close>
lemma Lset_in_Lset_succ: "Lset(i) \<in> Lset(succ(i))"
apply (simp add: Lset_succ DPow_def)
apply (rule_tac x=Nil in bexI)
@@ -906,7 +906,7 @@
apply (blast intro!: le_imp_subset Lset_mono)
done
-text{*Kunen's VI 1.11 *}
+text\<open>Kunen's VI 1.11\<close>
lemma Lset_subset_Vset: "Ord(i) ==> Lset(i) \<subseteq> Vset(i)"
apply (erule trans_induct)
apply (subst Lset)
@@ -916,14 +916,14 @@
apply (rule Pow_mono, blast)
done
-text{*Kunen's VI 1.12 *}
+text\<open>Kunen's VI 1.12\<close>
lemma Lset_subset_Vset': "i \<in> nat ==> Lset(i) = Vset(i)"
apply (erule nat_induct)
apply (simp add: Vfrom_0)
apply (simp add: Lset_succ Vset_succ Finite_Vset Finite_DPow_eq_Pow)
done
-text{*Every set of constructible sets is included in some @{term Lset}*}
+text\<open>Every set of constructible sets is included in some @{term Lset}\<close>
lemma subset_Lset:
"(\<forall>x\<in>A. L(x)) ==> \<exists>i. Ord(i) & A \<subseteq> Lset(i)"
by (rule_tac x = "\<Union>x\<in>A. succ(lrank(x))" in exI, force)
@@ -934,7 +934,7 @@
==> P"
by (blast dest: subset_Lset)
-subsubsection{*For L to satisfy the Powerset axiom *}
+subsubsection\<open>For L to satisfy the Powerset axiom\<close>
lemma LPow_env_typing:
"[| y \<in> Lset(i); Ord(i); y \<subseteq> X |]
@@ -949,7 +949,7 @@
apply (simp add: DPow_def)
apply (intro conjI, clarify)
apply (rule_tac a=x in UN_I, simp+)
-txt{*Now to create the formula @{term "y \<subseteq> X"} *}
+txt\<open>Now to create the formula @{term "y \<subseteq> X"}\<close>
apply (rule_tac x="Cons(X,Nil)" in bexI)
apply (rule_tac x="subset_fm(0,1)" in bexI)
apply typecheck
@@ -964,7 +964,7 @@
by (blast intro: L_I dest: L_D LPow_in_Lset)
-subsection{*Eliminating @{term arity} from the Definition of @{term Lset}*}
+subsection\<open>Eliminating @{term arity} from the Definition of @{term Lset}\<close>
lemma nth_zero_eq_0: "n \<in> nat ==> nth(n,[0]) = 0"
by (induct_tac n, auto)
@@ -995,7 +995,7 @@
done
-text{*A simpler version of @{term DPow}: no arity check!*}
+text\<open>A simpler version of @{term DPow}: no arity check!\<close>
definition
DPow' :: "i => i" where
"DPow'(A) == {X \<in> Pow(A).
@@ -1022,8 +1022,8 @@
apply (erule DPow'_subset_DPow)
done
-text{*And thus we can relativize @{term Lset} without bothering with
- @{term arity} and @{term length}*}
+text\<open>And thus we can relativize @{term Lset} without bothering with
+ @{term arity} and @{term length}\<close>
lemma Lset_eq_transrec_DPow': "Lset(i) = transrec(i, %x f. \<Union>y\<in>x. DPow'(f`y))"
apply (rule_tac a=i in eps_induct)
apply (subst Lset)
@@ -1031,8 +1031,8 @@
apply (simp only: DPow_eq_DPow' [OF Transset_Lset], simp)
done
-text{*With this rule we can specify @{term p} later and don't worry about
- arities at all!*}
+text\<open>With this rule we can specify @{term p} later and don't worry about
+ arities at all!\<close>
lemma DPow_LsetI [rule_format]:
"[|\<forall>x\<in>Lset(i). P(x) \<longleftrightarrow> sats(Lset(i), p, Cons(x,env));
env \<in> list(Lset(i)); p \<in> formula|]
--- a/src/ZF/Constructible/Internalize.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/ZF/Constructible/Internalize.thy Thu Jul 23 14:25:05 2015 +0200
@@ -4,9 +4,9 @@
theory Internalize imports L_axioms Datatype_absolute begin
-subsection{*Internalized Forms of Data Structuring Operators*}
+subsection\<open>Internalized Forms of Data Structuring Operators\<close>
-subsubsection{*The Formula @{term is_Inl}, Internalized*}
+subsubsection\<open>The Formula @{term is_Inl}, Internalized\<close>
(* is_Inl(M,a,z) == \<exists>zero[M]. empty(M,zero) & pair(M,zero,a,z) *)
definition
@@ -36,7 +36,7 @@
done
-subsubsection{*The Formula @{term is_Inr}, Internalized*}
+subsubsection\<open>The Formula @{term is_Inr}, Internalized\<close>
(* is_Inr(M,a,z) == \<exists>n1[M]. number1(M,n1) & pair(M,n1,a,z) *)
definition
@@ -66,7 +66,7 @@
done
-subsubsection{*The Formula @{term is_Nil}, Internalized*}
+subsubsection\<open>The Formula @{term is_Nil}, Internalized\<close>
(* is_Nil(M,xs) == \<exists>zero[M]. empty(M,zero) & is_Inl(M,zero,xs) *)
@@ -95,7 +95,7 @@
done
-subsubsection{*The Formula @{term is_Cons}, Internalized*}
+subsubsection\<open>The Formula @{term is_Cons}, Internalized\<close>
(* "is_Cons(M,a,l,Z) == \<exists>p[M]. pair(M,a,l,p) & is_Inr(M,p,Z)" *)
@@ -127,7 +127,7 @@
apply (intro FOL_reflections pair_reflection Inr_reflection)
done
-subsubsection{*The Formula @{term is_quasilist}, Internalized*}
+subsubsection\<open>The Formula @{term is_quasilist}, Internalized\<close>
(* is_quasilist(M,xs) == is_Nil(M,z) | (\<exists>x[M]. \<exists>l[M]. is_Cons(M,x,l,z))" *)
@@ -157,10 +157,10 @@
done
-subsection{*Absoluteness for the Function @{term nth}*}
+subsection\<open>Absoluteness for the Function @{term nth}\<close>
-subsubsection{*The Formula @{term is_hd}, Internalized*}
+subsubsection\<open>The Formula @{term is_hd}, Internalized\<close>
(* "is_hd(M,xs,H) ==
(is_Nil(M,xs) \<longrightarrow> empty(M,H)) &
@@ -197,7 +197,7 @@
done
-subsubsection{*The Formula @{term is_tl}, Internalized*}
+subsubsection\<open>The Formula @{term is_tl}, Internalized\<close>
(* "is_tl(M,xs,T) ==
(is_Nil(M,xs) \<longrightarrow> T=xs) &
@@ -234,12 +234,12 @@
done
-subsubsection{*The Operator @{term is_bool_of_o}*}
+subsubsection\<open>The Operator @{term is_bool_of_o}\<close>
(* is_bool_of_o :: "[i=>o, o, i] => o"
"is_bool_of_o(M,P,z) == (P & number1(M,z)) | (~P & empty(M,z))" *)
-text{*The formula @{term p} has no free variables.*}
+text\<open>The formula @{term p} has no free variables.\<close>
definition
bool_of_o_fm :: "[i, i]=>i" where
"bool_of_o_fm(p,z) ==
@@ -272,12 +272,12 @@
done
-subsection{*More Internalizations*}
+subsection\<open>More Internalizations\<close>
-subsubsection{*The Operator @{term is_lambda}*}
+subsubsection\<open>The Operator @{term is_lambda}\<close>
-text{*The two arguments of @{term p} are always 1, 0. Remember that
- @{term p} will be enclosed by three quantifiers.*}
+text\<open>The two arguments of @{term p} are always 1, 0. Remember that
+ @{term p} will be enclosed by three quantifiers.\<close>
(* is_lambda :: "[i=>o, i, [i,i]=>o, i] => o"
"is_lambda(M, A, is_b, z) ==
@@ -290,8 +290,8 @@
Exists(Exists(And(Member(1,A#+3),
And(pair_fm(1,0,2), p))))))"
-text{*We call @{term p} with arguments x, y by equating them with
- the corresponding quantified variables with de Bruijn indices 1, 0.*}
+text\<open>We call @{term p} with arguments x, y by equating them with
+ the corresponding quantified variables with de Bruijn indices 1, 0.\<close>
lemma is_lambda_type [TC]:
"[| p \<in> formula; x \<in> nat; y \<in> nat |]
@@ -319,7 +319,7 @@
apply (intro FOL_reflections is_b_reflection pair_reflection)
done
-subsubsection{*The Operator @{term is_Member}, Internalized*}
+subsubsection\<open>The Operator @{term is_Member}, Internalized\<close>
(* "is_Member(M,x,y,Z) ==
\<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) & is_Inl(M,p,u) & is_Inl(M,u,Z)" *)
@@ -352,7 +352,7 @@
apply (intro FOL_reflections pair_reflection Inl_reflection)
done
-subsubsection{*The Operator @{term is_Equal}, Internalized*}
+subsubsection\<open>The Operator @{term is_Equal}, Internalized\<close>
(* "is_Equal(M,x,y,Z) ==
\<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) & is_Inr(M,p,u) & is_Inl(M,u,Z)" *)
@@ -385,7 +385,7 @@
apply (intro FOL_reflections pair_reflection Inl_reflection Inr_reflection)
done
-subsubsection{*The Operator @{term is_Nand}, Internalized*}
+subsubsection\<open>The Operator @{term is_Nand}, Internalized\<close>
(* "is_Nand(M,x,y,Z) ==
\<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) & is_Inl(M,p,u) & is_Inr(M,u,Z)" *)
@@ -418,7 +418,7 @@
apply (intro FOL_reflections pair_reflection Inl_reflection Inr_reflection)
done
-subsubsection{*The Operator @{term is_Forall}, Internalized*}
+subsubsection\<open>The Operator @{term is_Forall}, Internalized\<close>
(* "is_Forall(M,p,Z) == \<exists>u[M]. is_Inr(M,p,u) & is_Inr(M,u,Z)" *)
definition
@@ -450,7 +450,7 @@
done
-subsubsection{*The Operator @{term is_and}, Internalized*}
+subsubsection\<open>The Operator @{term is_and}, Internalized\<close>
(* is_and(M,a,b,z) == (number1(M,a) & z=b) |
(~number1(M,a) & empty(M,z)) *)
@@ -484,7 +484,7 @@
done
-subsubsection{*The Operator @{term is_or}, Internalized*}
+subsubsection\<open>The Operator @{term is_or}, Internalized\<close>
(* is_or(M,a,b,z) == (number1(M,a) & number1(M,z)) |
(~number1(M,a) & z=b) *)
@@ -520,7 +520,7 @@
-subsubsection{*The Operator @{term is_not}, Internalized*}
+subsubsection\<open>The Operator @{term is_not}, Internalized\<close>
(* is_not(M,a,z) == (number1(M,a) & empty(M,z)) |
(~number1(M,a) & number1(M,z)) *)
@@ -559,11 +559,11 @@
is_lambda_reflection Member_reflection Equal_reflection Nand_reflection
Forall_reflection is_and_reflection is_or_reflection is_not_reflection
-subsection{*Well-Founded Recursion!*}
+subsection\<open>Well-Founded Recursion!\<close>
-subsubsection{*The Operator @{term M_is_recfun}*}
+subsubsection\<open>The Operator @{term M_is_recfun}\<close>
-text{*Alternative definition, minimizing nesting of quantifiers around MH*}
+text\<open>Alternative definition, minimizing nesting of quantifiers around MH\<close>
lemma M_is_recfun_iff:
"M_is_recfun(M,MH,r,a,f) \<longleftrightarrow>
(\<forall>z[M]. z \<in> f \<longleftrightarrow>
@@ -590,7 +590,7 @@
xa \<in> r)"
*)
-text{*The three arguments of @{term p} are always 2, 1, 0 and z*}
+text\<open>The three arguments of @{term p} are always 2, 1, 0 and z\<close>
definition
is_recfun_fm :: "[i, i, i, i]=>i" where
"is_recfun_fm(p,r,a,f) ==
@@ -632,9 +632,9 @@
==> M_is_recfun(##A, MH, x, y, z) \<longleftrightarrow> sats(A, is_recfun_fm(p,i,j,k), env)"
by (simp add: sats_is_recfun_fm [OF MH_iff_sats])
-text{*The additional variable in the premise, namely @{term f'}, is essential.
+text\<open>The additional variable in the premise, namely @{term f'}, is essential.
It lets @{term MH} depend upon @{term x}, which seems often necessary.
-The same thing occurs in @{text is_wfrec_reflection}.*}
+The same thing occurs in @{text is_wfrec_reflection}.\<close>
theorem is_recfun_reflection:
assumes MH_reflection:
"!!f' f g h. REFLECTS[\<lambda>x. MH(L, f'(x), f(x), g(x), h(x)),
@@ -646,10 +646,10 @@
restriction_reflection MH_reflection)
done
-subsubsection{*The Operator @{term is_wfrec}*}
+subsubsection\<open>The Operator @{term is_wfrec}\<close>
-text{*The three arguments of @{term p} are always 2, 1, 0;
- @{term p} is enclosed by 5 quantifiers.*}
+text\<open>The three arguments of @{term p} are always 2, 1, 0;
+ @{term p} is enclosed by 5 quantifiers.\<close>
(* is_wfrec :: "[i=>o, i, [i,i,i]=>o, i, i] => o"
"is_wfrec(M,MH,r,a,z) ==
@@ -661,11 +661,11 @@
Exists(Exists(Exists(Exists(
And(Equal(2,a#+5), And(Equal(1,4), And(Equal(0,z#+5), p)))))))))"
-text{*We call @{term p} with arguments a, f, z by equating them with
- the corresponding quantified variables with de Bruijn indices 2, 1, 0.*}
+text\<open>We call @{term p} with arguments a, f, z by equating them with
+ the corresponding quantified variables with de Bruijn indices 2, 1, 0.\<close>
-text{*There's an additional existential quantifier to ensure that the
- environments in both calls to MH have the same length.*}
+text\<open>There's an additional existential quantifier to ensure that the
+ environments in both calls to MH have the same length.\<close>
lemma is_wfrec_type [TC]:
"[| p \<in> formula; x \<in> nat; y \<in> nat; z \<in> nat |]
@@ -709,9 +709,9 @@
done
-subsection{*For Datatypes*}
+subsection\<open>For Datatypes\<close>
-subsubsection{*Binary Products, Internalized*}
+subsubsection\<open>Binary Products, Internalized\<close>
definition
cartprod_fm :: "[i,i,i]=>i" where
@@ -747,7 +747,7 @@
done
-subsubsection{*Binary Sums, Internalized*}
+subsubsection\<open>Binary Sums, Internalized\<close>
(* "is_sum(M,A,B,Z) ==
\<exists>A0[M]. \<exists>n1[M]. \<exists>s1[M]. \<exists>B1[M].
@@ -787,7 +787,7 @@
done
-subsubsection{*The Operator @{term quasinat}*}
+subsubsection\<open>The Operator @{term quasinat}\<close>
(* "is_quasinat(M,z) == empty(M,z) | (\<exists>m[M]. successor(M,m,z))" *)
definition
@@ -817,17 +817,17 @@
done
-subsubsection{*The Operator @{term is_nat_case}*}
-text{*I could not get it to work with the more natural assumption that
+subsubsection\<open>The Operator @{term is_nat_case}\<close>
+text\<open>I could not get it to work with the more natural assumption that
@{term is_b} takes two arguments. Instead it must be a formula where 1 and 0
- stand for @{term m} and @{term b}, respectively.*}
+ stand for @{term m} and @{term b}, respectively.\<close>
(* is_nat_case :: "[i=>o, i, [i,i]=>o, i, i] => o"
"is_nat_case(M, a, is_b, k, z) ==
(empty(M,k) \<longrightarrow> z=a) &
(\<forall>m[M]. successor(M,m,k) \<longrightarrow> is_b(m,z)) &
(is_quasinat(M,k) | empty(M,z))" *)
-text{*The formula @{term is_b} has free variables 1 and 0.*}
+text\<open>The formula @{term is_b} has free variables 1 and 0.\<close>
definition
is_nat_case_fm :: "[i, i, i, i]=>i" where
"is_nat_case_fm(a,is_b,k,z) ==
@@ -863,9 +863,9 @@
by (simp add: sats_is_nat_case_fm [of A is_b])
-text{*The second argument of @{term is_b} gives it direct access to @{term x},
+text\<open>The second argument of @{term is_b} gives it direct access to @{term x},
which is essential for handling free variable references. Without this
- argument, we cannot prove reflection for @{term iterates_MH}.*}
+ argument, we cannot prove reflection for @{term iterates_MH}.\<close>
theorem is_nat_case_reflection:
assumes is_b_reflection:
"!!h f g. REFLECTS[\<lambda>x. is_b(L, h(x), f(x), g(x)),
@@ -878,7 +878,7 @@
done
-subsection{*The Operator @{term iterates_MH}, Needed for Iteration*}
+subsection\<open>The Operator @{term iterates_MH}, Needed for Iteration\<close>
(* iterates_MH :: "[i=>o, [i,i]=>o, i, i, i, i] => o"
"iterates_MH(M,isF,v,n,g,z) ==
@@ -926,9 +926,9 @@
sats(A, iterates_MH_fm(p,i',i,j,k), env)"
by (simp add: sats_iterates_MH_fm [OF is_F_iff_sats])
-text{*The second argument of @{term p} gives it direct access to @{term x},
+text\<open>The second argument of @{term p} gives it direct access to @{term x},
which is essential for handling free variable references. Without this
- argument, we cannot prove reflection for @{term list_N}.*}
+ argument, we cannot prove reflection for @{term list_N}.\<close>
theorem iterates_MH_reflection:
assumes p_reflection:
"!!f g h. REFLECTS[\<lambda>x. p(L, h(x), f(x), g(x)),
@@ -941,10 +941,10 @@
done
-subsubsection{*The Operator @{term is_iterates}*}
+subsubsection\<open>The Operator @{term is_iterates}\<close>
-text{*The three arguments of @{term p} are always 2, 1, 0;
- @{term p} is enclosed by 9 (??) quantifiers.*}
+text\<open>The three arguments of @{term p} are always 2, 1, 0;
+ @{term p} is enclosed by 9 (??) quantifiers.\<close>
(* "is_iterates(M,isF,v,n,Z) ==
\<exists>sn[M]. \<exists>msn[M]. successor(M,n,sn) & membership(M,sn,msn) &
@@ -959,8 +959,8 @@
is_wfrec_fm(iterates_MH_fm(p, v#+7, 2, 1, 0),
0, n#+2, Z#+2)))))"
-text{*We call @{term p} with arguments a, f, z by equating them with
- the corresponding quantified variables with de Bruijn indices 2, 1, 0.*}
+text\<open>We call @{term p} with arguments a, f, z by equating them with
+ the corresponding quantified variables with de Bruijn indices 2, 1, 0.\<close>
lemma is_iterates_type [TC]:
@@ -1002,9 +1002,9 @@
sats(A, is_iterates_fm(p,i,j,k), env)"
by (simp add: sats_is_iterates_fm [OF is_F_iff_sats])
-text{*The second argument of @{term p} gives it direct access to @{term x},
+text\<open>The second argument of @{term p} gives it direct access to @{term x},
which is essential for handling free variable references. Without this
- argument, we cannot prove reflection for @{term list_N}.*}
+ argument, we cannot prove reflection for @{term list_N}.\<close>
theorem is_iterates_reflection:
assumes p_reflection:
"!!f g h. REFLECTS[\<lambda>x. p(L, h(x), f(x), g(x)),
@@ -1017,7 +1017,7 @@
done
-subsubsection{*The Formula @{term is_eclose_n}, Internalized*}
+subsubsection\<open>The Formula @{term is_eclose_n}, Internalized\<close>
(* is_eclose_n(M,A,n,Z) == is_iterates(M, big_union(M), A, n, Z) *)
@@ -1053,7 +1053,7 @@
done
-subsubsection{*Membership in @{term "eclose(A)"}*}
+subsubsection\<open>Membership in @{term "eclose(A)"}\<close>
(* mem_eclose(M,A,l) ==
\<exists>n[M]. \<exists>eclosen[M].
@@ -1088,7 +1088,7 @@
done
-subsubsection{*The Predicate ``Is @{term "eclose(A)"}''*}
+subsubsection\<open>The Predicate ``Is @{term "eclose(A)"}''\<close>
(* is_eclose(M,A,Z) == \<forall>l[M]. l \<in> Z \<longleftrightarrow> mem_eclose(M,A,l) *)
definition
@@ -1119,7 +1119,7 @@
done
-subsubsection{*The List Functor, Internalized*}
+subsubsection\<open>The List Functor, Internalized\<close>
definition
list_functor_fm :: "[i,i,i]=>i" where
@@ -1156,7 +1156,7 @@
done
-subsubsection{*The Formula @{term is_list_N}, Internalized*}
+subsubsection\<open>The Formula @{term is_list_N}, Internalized\<close>
(* "is_list_N(M,A,n,Z) ==
\<exists>zero[M]. empty(M,zero) &
@@ -1198,7 +1198,7 @@
-subsubsection{*The Predicate ``Is A List''*}
+subsubsection\<open>The Predicate ``Is A List''\<close>
(* mem_list(M,A,l) ==
\<exists>n[M]. \<exists>listn[M].
@@ -1233,7 +1233,7 @@
done
-subsubsection{*The Predicate ``Is @{term "list(A)"}''*}
+subsubsection\<open>The Predicate ``Is @{term "list(A)"}''\<close>
(* is_list(M,A,Z) == \<forall>l[M]. l \<in> Z \<longleftrightarrow> mem_list(M,A,l) *)
definition
@@ -1264,7 +1264,7 @@
done
-subsubsection{*The Formula Functor, Internalized*}
+subsubsection\<open>The Formula Functor, Internalized\<close>
definition formula_functor_fm :: "[i,i]=>i" where
(* "is_formula_functor(M,X,Z) ==
@@ -1307,7 +1307,7 @@
done
-subsubsection{*The Formula @{term is_formula_N}, Internalized*}
+subsubsection\<open>The Formula @{term is_formula_N}, Internalized\<close>
(* "is_formula_N(M,n,Z) ==
\<exists>zero[M]. empty(M,zero) &
@@ -1348,7 +1348,7 @@
-subsubsection{*The Predicate ``Is A Formula''*}
+subsubsection\<open>The Predicate ``Is A Formula''\<close>
(* mem_formula(M,p) ==
\<exists>n[M]. \<exists>formn[M].
@@ -1383,7 +1383,7 @@
-subsubsection{*The Predicate ``Is @{term "formula"}''*}
+subsubsection\<open>The Predicate ``Is @{term "formula"}''\<close>
(* is_formula(M,Z) == \<forall>p[M]. p \<in> Z \<longleftrightarrow> mem_formula(M,p) *)
definition
@@ -1412,12 +1412,12 @@
done
-subsubsection{*The Operator @{term is_transrec}*}
+subsubsection\<open>The Operator @{term is_transrec}\<close>
-text{*The three arguments of @{term p} are always 2, 1, 0. It is buried
+text\<open>The three arguments of @{term p} are always 2, 1, 0. It is buried
within eight quantifiers!
We call @{term p} with arguments a, f, z by equating them with
- the corresponding quantified variables with de Bruijn indices 2, 1, 0.*}
+ the corresponding quantified variables with de Bruijn indices 2, 1, 0.\<close>
(* is_transrec :: "[i=>o, [i,i,i]=>o, i, i] => o"
"is_transrec(M,MH,a,z) ==
--- a/src/ZF/Constructible/L_axioms.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/ZF/Constructible/L_axioms.thy Thu Jul 23 14:25:05 2015 +0200
@@ -2,11 +2,11 @@
Author: Lawrence C Paulson, Cambridge University Computer Laboratory
*)
-section {* The ZF Axioms (Except Separation) in L *}
+section \<open>The ZF Axioms (Except Separation) in L\<close>
theory L_axioms imports Formula Relative Reflection MetaExists begin
-text {* The class L satisfies the premises of locale @{text M_trivial} *}
+text \<open>The class L satisfies the premises of locale @{text M_trivial}\<close>
lemma transL: "[| y\<in>x; L(x) |] ==> L(y)"
apply (insert Transset_Lset)
@@ -38,7 +38,7 @@
apply (blast intro: transL)
done
-text{*We don't actually need @{term L} to satisfy the foundation axiom.*}
+text\<open>We don't actually need @{term L} to satisfy the foundation axiom.\<close>
theorem foundation_ax: "foundation_ax(L)"
apply (simp add: foundation_ax_def)
apply (rule rallI)
@@ -46,7 +46,7 @@
apply (blast intro: transL)
done
-subsection{*For L to satisfy Replacement *}
+subsection\<open>For L to satisfy Replacement\<close>
(*Can't move these to Formula unless the definition of univalent is moved
there too!*)
@@ -78,8 +78,8 @@
apply (simp_all add: Replace_iff univalent_def, blast)
done
-subsection{*Instantiating the locale @{text M_trivial}*}
-text{*No instances of Separation yet.*}
+subsection\<open>Instantiating the locale @{text M_trivial}\<close>
+text\<open>No instances of Separation yet.\<close>
lemma Lset_mono_le: "mono_le_subset(Lset)"
by (simp add: mono_le_subset_def le_imp_subset Lset_mono)
@@ -110,9 +110,9 @@
...and dozens of similar ones.
*)
-subsection{*Instantiation of the locale @{text reflection}*}
+subsection\<open>Instantiation of the locale @{text reflection}\<close>
-text{*instances of locale constants*}
+text\<open>instances of locale constants\<close>
definition
L_F0 :: "[i=>o,i] => i" where
@@ -127,8 +127,8 @@
"L_ClEx(P) == \<lambda>a. Limit(a) \<and> normalize(L_FF(P),a) = a"
-text{*We must use the meta-existential quantifier; otherwise the reflection
- terms become enormous!*}
+text\<open>We must use the meta-existential quantifier; otherwise the reflection
+ terms become enormous!\<close>
definition
L_Reflects :: "[i=>o,[i,i]=>o] => prop" ("(3REFLECTS/ [_,/ _])") where
"REFLECTS[P,Q] == (??Cl. Closed_Unbounded(Cl) &
@@ -223,8 +223,8 @@
apply (intro Imp_reflection All_reflection, assumption)
done
-text{*This version handles an alternative form of the bounded quantifier
- in the second argument of @{text REFLECTS}.*}
+text\<open>This version handles an alternative form of the bounded quantifier
+ in the second argument of @{text REFLECTS}.\<close>
theorem Rex_reflection':
"REFLECTS[\<lambda>x. P(fst(x),snd(x)), \<lambda>a x. Q(a,fst(x),snd(x))]
==> REFLECTS[\<lambda>x. \<exists>z[L]. P(x,z), \<lambda>a x. \<exists>z[##Lset(a)]. Q(a,x,z)]"
@@ -232,7 +232,7 @@
apply (erule Rex_reflection [unfolded rex_def Bex_def])
done
-text{*As above.*}
+text\<open>As above.\<close>
theorem Rall_reflection':
"REFLECTS[\<lambda>x. P(fst(x),snd(x)), \<lambda>a x. Q(a,fst(x),snd(x))]
==> REFLECTS[\<lambda>x. \<forall>z[L]. P(x,z), \<lambda>a x. \<forall>z[##Lset(a)]. Q(a,x,z)]"
@@ -263,9 +263,9 @@
by blast
-subsection{*Internalized Formulas for some Set-Theoretic Concepts*}
+subsection\<open>Internalized Formulas for some Set-Theoretic Concepts\<close>
-subsubsection{*Some numbers to help write de Bruijn indices*}
+subsubsection\<open>Some numbers to help write de Bruijn indices\<close>
abbreviation
digit3 :: i ("3") where "3 == succ(2)"
@@ -289,7 +289,7 @@
digit9 :: i ("9") where "9 == succ(8)"
-subsubsection{*The Empty Set, Internalized*}
+subsubsection\<open>The Empty Set, Internalized\<close>
definition
empty_fm :: "i=>i" where
@@ -317,7 +317,7 @@
apply (intro FOL_reflections)
done
-text{*Not used. But maybe useful?*}
+text\<open>Not used. But maybe useful?\<close>
lemma Transset_sats_empty_fm_eq_0:
"[| n \<in> nat; env \<in> list(A); Transset(A)|]
==> sats(A, empty_fm(n), env) \<longleftrightarrow> nth(n,env) = 0"
@@ -328,7 +328,7 @@
done
-subsubsection{*Unordered Pairs, Internalized*}
+subsubsection\<open>Unordered Pairs, Internalized\<close>
definition
upair_fm :: "[i,i,i]=>i" where
@@ -354,7 +354,7 @@
==> upair(##A, x, y, z) \<longleftrightarrow> sats(A, upair_fm(i,j,k), env)"
by (simp add: sats_upair_fm)
-text{*Useful? At least it refers to "real" unordered pairs*}
+text\<open>Useful? At least it refers to "real" unordered pairs\<close>
lemma sats_upair_fm2 [simp]:
"[| x \<in> nat; y \<in> nat; z < length(env); env \<in> list(A); Transset(A)|]
==> sats(A, upair_fm(x,y,z), env) \<longleftrightarrow>
@@ -371,7 +371,7 @@
apply (intro FOL_reflections)
done
-subsubsection{*Ordered pairs, Internalized*}
+subsubsection\<open>Ordered pairs, Internalized\<close>
definition
pair_fm :: "[i,i,i]=>i" where
@@ -404,7 +404,7 @@
done
-subsubsection{*Binary Unions, Internalized*}
+subsubsection\<open>Binary Unions, Internalized\<close>
definition
union_fm :: "[i,i,i]=>i" where
@@ -436,7 +436,7 @@
done
-subsubsection{*Set ``Cons,'' Internalized*}
+subsubsection\<open>Set ``Cons,'' Internalized\<close>
definition
cons_fm :: "[i,i,i]=>i" where
@@ -469,7 +469,7 @@
done
-subsubsection{*Successor Function, Internalized*}
+subsubsection\<open>Successor Function, Internalized\<close>
definition
succ_fm :: "[i,i]=>i" where
@@ -499,7 +499,7 @@
done
-subsubsection{*The Number 1, Internalized*}
+subsubsection\<open>The Number 1, Internalized\<close>
(* "number1(M,a) == (\<exists>x[M]. empty(M,x) & successor(M,x,a))" *)
definition
@@ -529,7 +529,7 @@
done
-subsubsection{*Big Union, Internalized*}
+subsubsection\<open>Big Union, Internalized\<close>
(* "big_union(M,A,z) == \<forall>x[M]. x \<in> z \<longleftrightarrow> (\<exists>y[M]. y\<in>A & x \<in> y)" *)
definition
@@ -562,13 +562,13 @@
done
-subsubsection{*Variants of Satisfaction Definitions for Ordinals, etc.*}
+subsubsection\<open>Variants of Satisfaction Definitions for Ordinals, etc.\<close>
-text{*The @{text sats} theorems below are standard versions of the ones proved
+text\<open>The @{text sats} theorems below are standard versions of the ones proved
in theory @{text Formula}. They relate elements of type @{term formula} to
relativized concepts such as @{term subset} or @{term ordinal} rather than to
real concepts such as @{term Ord}. Now that we have instantiated the locale
-@{text M_trivial}, we no longer require the earlier versions.*}
+@{text M_trivial}, we no longer require the earlier versions.\<close>
lemma sats_subset_fm':
"[|x \<in> nat; y \<in> nat; env \<in> list(A)|]
@@ -611,7 +611,7 @@
done
-subsubsection{*Membership Relation, Internalized*}
+subsubsection\<open>Membership Relation, Internalized\<close>
definition
Memrel_fm :: "[i,i]=>i" where
@@ -645,7 +645,7 @@
apply (intro FOL_reflections pair_reflection)
done
-subsubsection{*Predecessor Set, Internalized*}
+subsubsection\<open>Predecessor Set, Internalized\<close>
definition
pred_set_fm :: "[i,i,i,i]=>i" where
@@ -682,7 +682,7 @@
-subsubsection{*Domain of a Relation, Internalized*}
+subsubsection\<open>Domain of a Relation, Internalized\<close>
(* "is_domain(M,r,z) ==
\<forall>x[M]. (x \<in> z \<longleftrightarrow> (\<exists>w[M]. w\<in>r & (\<exists>y[M]. pair(M,x,y,w))))" *)
@@ -717,7 +717,7 @@
done
-subsubsection{*Range of a Relation, Internalized*}
+subsubsection\<open>Range of a Relation, Internalized\<close>
(* "is_range(M,r,z) ==
\<forall>y[M]. (y \<in> z \<longleftrightarrow> (\<exists>w[M]. w\<in>r & (\<exists>x[M]. pair(M,x,y,w))))" *)
@@ -752,7 +752,7 @@
done
-subsubsection{*Field of a Relation, Internalized*}
+subsubsection\<open>Field of a Relation, Internalized\<close>
(* "is_field(M,r,z) ==
\<exists>dr[M]. is_domain(M,r,dr) &
@@ -789,7 +789,7 @@
done
-subsubsection{*Image under a Relation, Internalized*}
+subsubsection\<open>Image under a Relation, Internalized\<close>
(* "image(M,r,A,z) ==
\<forall>y[M]. (y \<in> z \<longleftrightarrow> (\<exists>w[M]. w\<in>r & (\<exists>x[M]. x\<in>A & pair(M,x,y,w))))" *)
@@ -825,7 +825,7 @@
done
-subsubsection{*Pre-Image under a Relation, Internalized*}
+subsubsection\<open>Pre-Image under a Relation, Internalized\<close>
(* "pre_image(M,r,A,z) ==
\<forall>x[M]. x \<in> z \<longleftrightarrow> (\<exists>w[M]. w\<in>r & (\<exists>y[M]. y\<in>A & pair(M,x,y,w)))" *)
@@ -861,7 +861,7 @@
done
-subsubsection{*Function Application, Internalized*}
+subsubsection\<open>Function Application, Internalized\<close>
(* "fun_apply(M,f,x,y) ==
(\<exists>xs[M]. \<exists>fxs[M].
@@ -898,7 +898,7 @@
done
-subsubsection{*The Concept of Relation, Internalized*}
+subsubsection\<open>The Concept of Relation, Internalized\<close>
(* "is_relation(M,r) ==
(\<forall>z[M]. z\<in>r \<longrightarrow> (\<exists>x[M]. \<exists>y[M]. pair(M,x,y,z)))" *)
@@ -930,7 +930,7 @@
done
-subsubsection{*The Concept of Function, Internalized*}
+subsubsection\<open>The Concept of Function, Internalized\<close>
(* "is_function(M,r) ==
\<forall>x[M]. \<forall>y[M]. \<forall>y'[M]. \<forall>p[M]. \<forall>p'[M].
@@ -967,7 +967,7 @@
done
-subsubsection{*Typed Functions, Internalized*}
+subsubsection\<open>Typed Functions, Internalized\<close>
(* "typed_function(M,A,B,r) ==
is_function(M,r) & is_relation(M,r) & is_domain(M,r,A) &
@@ -1026,7 +1026,7 @@
done
-subsubsection{*Composition of Relations, Internalized*}
+subsubsection\<open>Composition of Relations, Internalized\<close>
(* "composition(M,r,s,t) ==
\<forall>p[M]. p \<in> t \<longleftrightarrow>
@@ -1067,7 +1067,7 @@
done
-subsubsection{*Injections, Internalized*}
+subsubsection\<open>Injections, Internalized\<close>
(* "injection(M,A,B,f) ==
typed_function(M,A,B,f) &
@@ -1108,7 +1108,7 @@
done
-subsubsection{*Surjections, Internalized*}
+subsubsection\<open>Surjections, Internalized\<close>
(* surjection :: "[i=>o,i,i,i] => o"
"surjection(M,A,B,f) ==
@@ -1147,7 +1147,7 @@
-subsubsection{*Bijections, Internalized*}
+subsubsection\<open>Bijections, Internalized\<close>
(* bijection :: "[i=>o,i,i,i] => o"
"bijection(M,A,B,f) == injection(M,A,B,f) & surjection(M,A,B,f)" *)
@@ -1179,7 +1179,7 @@
done
-subsubsection{*Restriction of a Relation, Internalized*}
+subsubsection\<open>Restriction of a Relation, Internalized\<close>
(* "restriction(M,r,A,z) ==
@@ -1215,7 +1215,7 @@
apply (intro FOL_reflections pair_reflection)
done
-subsubsection{*Order-Isomorphisms, Internalized*}
+subsubsection\<open>Order-Isomorphisms, Internalized\<close>
(* order_isomorphism :: "[i=>o,i,i,i,i,i] => o"
"order_isomorphism(M,A,r,B,s,f) ==
@@ -1266,9 +1266,9 @@
apply (intro FOL_reflections function_reflections bijection_reflection)
done
-subsubsection{*Limit Ordinals, Internalized*}
+subsubsection\<open>Limit Ordinals, Internalized\<close>
-text{*A limit ordinal is a non-empty, successor-closed ordinal*}
+text\<open>A limit ordinal is a non-empty, successor-closed ordinal\<close>
(* "limit_ordinal(M,a) ==
ordinal(M,a) & ~ empty(M,a) &
@@ -1306,7 +1306,7 @@
empty_reflection successor_reflection)
done
-subsubsection{*Finite Ordinals: The Predicate ``Is A Natural Number''*}
+subsubsection\<open>Finite Ordinals: The Predicate ``Is A Natural Number''\<close>
(* "finite_ordinal(M,a) ==
ordinal(M,a) & ~ limit_ordinal(M,a) &
@@ -1342,7 +1342,7 @@
done
-subsubsection{*Omega: The Set of Natural Numbers*}
+subsubsection\<open>Omega: The Set of Natural Numbers\<close>
(* omega(M,a) == limit_ordinal(M,a) & (\<forall>x[M]. x\<in>a \<longrightarrow> ~ limit_ordinal(M,x)) *)
definition
--- a/src/ZF/Constructible/MetaExists.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/ZF/Constructible/MetaExists.thy Thu Jul 23 14:25:05 2015 +0200
@@ -2,12 +2,12 @@
Author: Lawrence C Paulson, Cambridge University Computer Laboratory
*)
-section{*The meta-existential quantifier*}
+section\<open>The meta-existential quantifier\<close>
theory MetaExists imports Main begin
-text{*Allows quantification over any term having sort @{text logic}. Used to
-quantify over classes. Yields a proposition rather than a FOL formula.*}
+text\<open>Allows quantification over any term having sort @{text logic}. Used to
+quantify over classes. Yields a proposition rather than a FOL formula.\<close>
definition
ex :: "(('a::{}) => prop) => prop" (binder "?? " 0) where
--- a/src/ZF/Constructible/Normal.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/ZF/Constructible/Normal.thy Thu Jul 23 14:25:05 2015 +0200
@@ -2,20 +2,20 @@
Author: Lawrence C Paulson, Cambridge University Computer Laboratory
*)
-section {*Closed Unbounded Classes and Normal Functions*}
+section \<open>Closed Unbounded Classes and Normal Functions\<close>
theory Normal imports Main begin
-text{*
+text\<open>
One source is the book
Frank R. Drake.
\emph{Set Theory: An Introduction to Large Cardinals}.
North-Holland, 1974.
-*}
+\<close>
-subsection {*Closed and Unbounded (c.u.) Classes of Ordinals*}
+subsection \<open>Closed and Unbounded (c.u.) Classes of Ordinals\<close>
definition
Closed :: "(i=>o) => o" where
@@ -30,7 +30,7 @@
"Closed_Unbounded(P) == Closed(P) \<and> Unbounded(P)"
-subsubsection{*Simple facts about c.u. classes*}
+subsubsection\<open>Simple facts about c.u. classes\<close>
lemma ClosedI:
"[| !!I. [| I \<noteq> 0; \<forall>i\<in>I. Ord(i) \<and> P(i) |] ==> P(\<Union>(I)) |]
@@ -50,16 +50,16 @@
by (simp add: Closed_Unbounded_def)
-text{*The universal class, V, is closed and unbounded.
- A bit odd, since C. U. concerns only ordinals, but it's used below!*}
+text\<open>The universal class, V, is closed and unbounded.
+ A bit odd, since C. U. concerns only ordinals, but it's used below!\<close>
theorem Closed_Unbounded_V [simp]: "Closed_Unbounded(\<lambda>x. True)"
by (unfold Closed_Unbounded_def Closed_def Unbounded_def, blast)
-text{*The class of ordinals, @{term Ord}, is closed and unbounded.*}
+text\<open>The class of ordinals, @{term Ord}, is closed and unbounded.\<close>
theorem Closed_Unbounded_Ord [simp]: "Closed_Unbounded(Ord)"
by (unfold Closed_Unbounded_def Closed_def Unbounded_def, blast)
-text{*The class of limit ordinals, @{term Limit}, is closed and unbounded.*}
+text\<open>The class of limit ordinals, @{term Limit}, is closed and unbounded.\<close>
theorem Closed_Unbounded_Limit [simp]: "Closed_Unbounded(Limit)"
apply (simp add: Closed_Unbounded_def Closed_def Unbounded_def Limit_Union,
clarify)
@@ -67,17 +67,17 @@
apply (blast intro: oadd_lt_self oadd_LimitI Limit_nat Limit_has_0)
done
-text{*The class of cardinals, @{term Card}, is closed and unbounded.*}
+text\<open>The class of cardinals, @{term Card}, is closed and unbounded.\<close>
theorem Closed_Unbounded_Card [simp]: "Closed_Unbounded(Card)"
apply (simp add: Closed_Unbounded_def Closed_def Unbounded_def Card_Union)
apply (blast intro: lt_csucc Card_csucc)
done
-subsubsection{*The intersection of any set-indexed family of c.u. classes is
- c.u.*}
+subsubsection\<open>The intersection of any set-indexed family of c.u. classes is
+ c.u.\<close>
-text{*The constructions below come from Kunen, \emph{Set Theory}, page 78.*}
+text\<open>The constructions below come from Kunen, \emph{Set Theory}, page 78.\<close>
locale cub_family =
fixes P and A
fixes next_greater -- "the next ordinal satisfying class @{term A}"
@@ -89,11 +89,11 @@
and "sup_greater(x) == \<Union>a\<in>A. next_greater(a,x)"
-text{*Trivial that the intersection is closed.*}
+text\<open>Trivial that the intersection is closed.\<close>
lemma (in cub_family) Closed_INT: "Closed(\<lambda>x. \<forall>i\<in>A. P(i,x))"
by (blast intro: ClosedI ClosedD [OF closed])
-text{*All remaining effort goes to show that the intersection is unbounded.*}
+text\<open>All remaining effort goes to show that the intersection is unbounded.\<close>
lemma (in cub_family) Ord_sup_greater:
"Ord(sup_greater(x))"
@@ -103,8 +103,8 @@
"Ord(next_greater(a,x))"
by (simp add: next_greater_def Ord_Least)
-text{*@{term next_greater} works as expected: it returns a larger value
-and one that belongs to class @{term "P(a)"}. *}
+text\<open>@{term next_greater} works as expected: it returns a larger value
+and one that belongs to class @{term "P(a)"}.\<close>
lemma (in cub_family) next_greater_lemma:
"[| Ord(x); a\<in>A |] ==> P(a, next_greater(a,x)) \<and> x < next_greater(a,x)"
apply (simp add: next_greater_def)
@@ -142,9 +142,9 @@
apply (rule le_anti_sym)
apply (rule le_implies_UN_le_UN)
apply (blast intro: leI next_greater_gt Ord_iterates Ord_sup_greater)
-txt{*Opposite bound:
+txt\<open>Opposite bound:
@{subgoals[display,indent=0,margin=65]}
-*}
+\<close>
apply (rule UN_least_le)
apply (blast intro: Ord_UN Ord_iterates Ord_sup_greater)
apply (rule_tac a="succ(n)" in UN_upper_le)
@@ -199,7 +199,7 @@
done
-subsection {*Normal Functions*}
+subsection \<open>Normal Functions\<close>
definition
mono_le_subset :: "(i=>i) => o" where
@@ -218,7 +218,7 @@
"Normal(F) == mono_Ord(F) \<and> cont_Ord(F)"
-subsubsection{*Immediate properties of the definitions*}
+subsubsection\<open>Immediate properties of the definitions\<close>
lemma NormalI:
"[|!!i j. i<j ==> F(i) < F(j); !!l. Limit(l) ==> F(l) = (\<Union>i<l. F(i))|]
@@ -267,9 +267,9 @@
qed
-subsubsection{*The class of fixedpoints is closed and unbounded*}
+subsubsection\<open>The class of fixedpoints is closed and unbounded\<close>
-text{*The proof is from Drake, pages 113--114.*}
+text\<open>The proof is from Drake, pages 113--114.\<close>
lemma mono_Ord_imp_le_subset: "mono_Ord(F) ==> mono_le_subset(F)"
apply (simp add: mono_le_subset_def, clarify)
@@ -278,26 +278,26 @@
apply (blast intro: lt_Ord2 mono_Ord_imp_Ord mono_Ord_imp_mono)
done
-text{*The following equation is taken for granted in any set theory text.*}
+text\<open>The following equation is taken for granted in any set theory text.\<close>
lemma cont_Ord_Union:
"[| cont_Ord(F); mono_le_subset(F); X=0 \<longrightarrow> F(0)=0; \<forall>x\<in>X. Ord(x) |]
==> F(\<Union>(X)) = (\<Union>y\<in>X. F(y))"
apply (frule Ord_set_cases)
apply (erule disjE, force)
apply (thin_tac "X=0 \<longrightarrow> Q" for Q, auto)
- txt{*The trival case of @{term "\<Union>X \<in> X"}*}
+ txt\<open>The trival case of @{term "\<Union>X \<in> X"}\<close>
apply (rule equalityI, blast intro: Ord_Union_eq_succD)
apply (simp add: mono_le_subset_def UN_subset_iff le_subset_iff)
apply (blast elim: equalityE)
-txt{*The limit case, @{term "Limit(\<Union>X)"}:
+txt\<open>The limit case, @{term "Limit(\<Union>X)"}:
@{subgoals[display,indent=0,margin=65]}
-*}
+\<close>
apply (simp add: OUN_Union_eq cont_Ord_def)
apply (rule equalityI)
-txt{*First inclusion:*}
+txt\<open>First inclusion:\<close>
apply (rule UN_least [OF OUN_least])
apply (simp add: mono_le_subset_def, blast intro: leI)
-txt{*Second inclusion:*}
+txt\<open>Second inclusion:\<close>
apply (rule UN_least)
apply (frule Union_upper_le, blast, blast intro: Ord_Union)
apply (erule leE, drule ltD, elim UnionE)
@@ -329,7 +329,7 @@
"[| n\<in>nat; Normal(F); Ord(x) |] ==> Ord(F^n (x))"
by (simp add: Ord_iterates)
-text{*THIS RESULT IS UNUSED*}
+text\<open>THIS RESULT IS UNUSED\<close>
lemma iterates_omega_Limit:
"[| Normal(F); x < F(x) |] ==> Limit(F^\<omega> (x))"
apply (frule lt_Ord)
@@ -353,9 +353,9 @@
apply (simp_all add: iterates_omega_triv [OF sym]) (*for subgoal 2*)
apply (simp add: iterates_omega_def Normal_Union)
apply (rule equalityI, force simp add: nat_succI)
-txt{*Opposite inclusion:
+txt\<open>Opposite inclusion:
@{subgoals[display,indent=0,margin=65]}
-*}
+\<close>
apply clarify
apply (rule UN_I, assumption)
apply (frule iterates_Normal_increasing, assumption, assumption, simp)
@@ -382,13 +382,13 @@
Normal_imp_fp_Unbounded)
-subsubsection{*Function @{text normalize}*}
+subsubsection\<open>Function @{text normalize}\<close>
-text{*Function @{text normalize} maps a function @{text F} to a
+text\<open>Function @{text normalize} maps a function @{text F} to a
normal function that bounds it above. The result is normal if and
only if @{text F} is continuous: succ is not bounded above by any
normal function, by @{thm [source] Normal_imp_fp_Unbounded}.
-*}
+\<close>
definition
normalize :: "[i=>i, i] => i" where
"normalize(F,a) == transrec2(a, F(0), \<lambda>x r. F(succ(x)) \<union> succ(r))"
@@ -451,9 +451,9 @@
qed
-subsection {*The Alephs*}
-text {*This is the well-known transfinite enumeration of the cardinal
-numbers.*}
+subsection \<open>The Alephs\<close>
+text \<open>This is the well-known transfinite enumeration of the cardinal
+numbers.\<close>
definition
Aleph :: "i => i" where
--- a/src/ZF/Constructible/Rank.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/ZF/Constructible/Rank.thy Thu Jul 23 14:25:05 2015 +0200
@@ -2,12 +2,12 @@
Author: Lawrence C Paulson, Cambridge University Computer Laboratory
*)
-section {*Absoluteness for Order Types, Rank Functions and Well-Founded
- Relations*}
+section \<open>Absoluteness for Order Types, Rank Functions and Well-Founded
+ Relations\<close>
theory Rank imports WF_absolute begin
-subsection {*Order Types: A Direct Construction by Replacement*}
+subsection \<open>Order Types: A Direct Construction by Replacement\<close>
locale M_ordertype = M_basic +
assumes well_ord_iso_separation:
@@ -15,7 +15,7 @@
==> separation (M, \<lambda>x. x\<in>A \<longrightarrow> (\<exists>y[M]. (\<exists>p[M].
fun_apply(M,f,x,y) & pair(M,y,x,p) & p \<in> r)))"
and obase_separation:
- --{*part of the order type formalization*}
+ --\<open>part of the order type formalization\<close>
"[| M(A); M(r) |]
==> separation(M, \<lambda>a. \<exists>x[M]. \<exists>g[M]. \<exists>mx[M]. \<exists>par[M].
ordinal(M,x) & membership(M,x,mx) & pred_set(M,A,a,r,par) &
@@ -34,8 +34,8 @@
pred_set(M,A,a,r,par) & order_isomorphism(M,par,r,x,mx,g))"
-text{*Inductive argument for Kunen's Lemma I 6.1, etc.
- Simple proof from Halmos, page 72*}
+text\<open>Inductive argument for Kunen's Lemma I 6.1, etc.
+ Simple proof from Halmos, page 72\<close>
lemma (in M_ordertype) wellordered_iso_subset_lemma:
"[| wellordered(M,A,r); f \<in> ord_iso(A,r, A',r); A'<= A; y \<in> A;
M(A); M(f); M(r) |] ==> ~ <f`y, y> \<in> r"
@@ -48,8 +48,8 @@
done
-text{*Kunen's Lemma I 6.1, page 14:
- there's no order-isomorphism to an initial segment of a well-ordering*}
+text\<open>Kunen's Lemma I 6.1, page 14:
+ there's no order-isomorphism to an initial segment of a well-ordering\<close>
lemma (in M_ordertype) wellordered_iso_predD:
"[| wellordered(M,A,r); f \<in> ord_iso(A, r, Order.pred(A,x,r), r);
M(A); M(f); M(r) |] ==> x \<notin> A"
@@ -75,7 +75,7 @@
done
-text{*Simple consequence of Lemma 6.1*}
+text\<open>Simple consequence of Lemma 6.1\<close>
lemma (in M_ordertype) wellordered_iso_pred_eq:
"[| wellordered(M,A,r);
f \<in> ord_iso(Order.pred(A,a,r), r, Order.pred(A,c,r), r);
@@ -89,11 +89,11 @@
done
-text{*Following Kunen's Theorem I 7.6, page 17. Note that this material is
-not required elsewhere.*}
+text\<open>Following Kunen's Theorem I 7.6, page 17. Note that this material is
+not required elsewhere.\<close>
-text{*Can't use @{text well_ord_iso_preserving} because it needs the
-strong premise @{term "well_ord(A,r)"}*}
+text\<open>Can't use @{text well_ord_iso_preserving} because it needs the
+strong premise @{term "well_ord(A,r)"}\<close>
lemma (in M_ordertype) ord_iso_pred_imp_lt:
"[| f \<in> ord_iso(Order.pred(A,x,r), r, i, Memrel(i));
g \<in> ord_iso(Order.pred(A,y,r), r, j, Memrel(j));
@@ -103,18 +103,18 @@
apply (frule wellordered_is_trans_on, assumption)
apply (frule_tac y=y in transM, assumption)
apply (rule_tac i=i and j=j in Ord_linear_lt, auto)
-txt{*case @{term "i=j"} yields a contradiction*}
+txt\<open>case @{term "i=j"} yields a contradiction\<close>
apply (rule_tac x1=x and A1="Order.pred(A,y,r)" in
wellordered_iso_predD [THEN notE])
apply (blast intro: wellordered_subset [OF _ pred_subset])
apply (simp add: trans_pred_pred_eq)
apply (blast intro: Ord_iso_implies_eq ord_iso_sym ord_iso_trans)
apply (simp_all add: pred_iff pred_closed converse_closed comp_closed)
-txt{*case @{term "j<i"} also yields a contradiction*}
+txt\<open>case @{term "j<i"} also yields a contradiction\<close>
apply (frule restrict_ord_iso2, assumption+)
apply (frule ord_iso_sym [THEN ord_iso_is_bij, THEN bij_is_fun])
apply (frule apply_type, blast intro: ltD)
- --{*thus @{term "converse(f)`j \<in> Order.pred(A,x,r)"}*}
+ --\<open>thus @{term "converse(f)`j \<in> Order.pred(A,x,r)"}\<close>
apply (simp add: pred_iff)
apply (subgoal_tac
"\<exists>h[M]. h \<in> ord_iso(Order.pred(A,y,r), r,
@@ -137,26 +137,26 @@
definition
obase :: "[i=>o,i,i] => i" where
- --{*the domain of @{text om}, eventually shown to equal @{text A}*}
+ --\<open>the domain of @{text om}, eventually shown to equal @{text A}\<close>
"obase(M,A,r) == {a\<in>A. \<exists>x[M]. \<exists>g[M]. Ord(x) &
g \<in> ord_iso(Order.pred(A,a,r),r,x,Memrel(x))}"
definition
omap :: "[i=>o,i,i,i] => o" where
- --{*the function that maps wosets to order types*}
+ --\<open>the function that maps wosets to order types\<close>
"omap(M,A,r,f) ==
\<forall>z[M].
z \<in> f \<longleftrightarrow> (\<exists>a\<in>A. \<exists>x[M]. \<exists>g[M]. z = <a,x> & Ord(x) &
g \<in> ord_iso(Order.pred(A,a,r),r,x,Memrel(x)))"
definition
- otype :: "[i=>o,i,i,i] => o" where --{*the order types themselves*}
+ otype :: "[i=>o,i,i,i] => o" where --\<open>the order types themselves\<close>
"otype(M,A,r,i) == \<exists>f[M]. omap(M,A,r,f) & is_range(M,f,i)"
-text{*Can also be proved with the premise @{term "M(z)"} instead of
+text\<open>Can also be proved with the premise @{term "M(z)"} instead of
@{term "M(f)"}, but that version is less useful. This lemma
- is also more useful than the definition, @{text omap_def}.*}
+ is also more useful than the definition, @{text omap_def}.\<close>
lemma (in M_ordertype) omap_iff:
"[| omap(M,A,r,f); M(A); M(f) |]
==> z \<in> f \<longleftrightarrow>
@@ -256,7 +256,7 @@
done
-text{*This is not the final result: we must show @{term "oB(A,r) = A"}*}
+text\<open>This is not the final result: we must show @{term "oB(A,r) = A"}\<close>
lemma (in M_ordertype) omap_ord_iso:
"[| wellordered(M,A,r); omap(M,A,r,f); otype(M,A,r,i);
M(A); M(r); M(f); M(i) |] ==> f \<in> ord_iso(obase(M,A,r),r,i,Memrel(i))"
@@ -266,15 +266,15 @@
apply (frule_tac a=x in apply_Pair, assumption)
apply (frule_tac a=y in apply_Pair, assumption)
apply (auto simp add: omap_iff)
- txt{*direction 1: assuming @{term "\<langle>x,y\<rangle> \<in> r"}*}
+ txt\<open>direction 1: assuming @{term "\<langle>x,y\<rangle> \<in> r"}\<close>
apply (blast intro: ltD ord_iso_pred_imp_lt)
- txt{*direction 2: proving @{term "\<langle>x,y\<rangle> \<in> r"} using linearity of @{term r}*}
+ txt\<open>direction 2: proving @{term "\<langle>x,y\<rangle> \<in> r"} using linearity of @{term r}\<close>
apply (rename_tac x y g ga)
apply (frule wellordered_is_linear, assumption,
erule_tac x=x and y=y in linearE, assumption+)
-txt{*the case @{term "x=y"} leads to immediate contradiction*}
+txt\<open>the case @{term "x=y"} leads to immediate contradiction\<close>
apply (blast elim: mem_irrefl)
-txt{*the case @{term "\<langle>y,x\<rangle> \<in> r"}: handle like the opposite direction*}
+txt\<open>the case @{term "\<langle>y,x\<rangle> \<in> r"}: handle like the opposite direction\<close>
apply (blast dest: ord_iso_pred_imp_lt ltD elim: mem_asym)
done
@@ -284,7 +284,7 @@
apply (frule wellordered_is_trans_on, assumption)
apply (rule OrdI)
prefer 2 apply (simp add: image_iff omap_iff Ord_def, blast)
-txt{*Hard part is to show that the image is a transitive set.*}
+txt\<open>Hard part is to show that the image is a transitive set.\<close>
apply (simp add: Transset_def, clarify)
apply (simp add: image_iff pred_iff apply_iff [OF omap_funtype [of A r f i]])
apply (rename_tac c j, clarify)
@@ -331,8 +331,8 @@
-text{*Main result: @{term om} gives the order-isomorphism
- @{term "\<langle>A,r\<rangle> \<cong> \<langle>i, Memrel(i)\<rangle>"} *}
+text\<open>Main result: @{term om} gives the order-isomorphism
+ @{term "\<langle>A,r\<rangle> \<cong> \<langle>i, Memrel(i)\<rangle>"}\<close>
theorem (in M_ordertype) omap_ord_iso_otype:
"[| wellordered(M,A,r); omap(M,A,r,f); otype(M,A,r,i);
M(A); M(r); M(f); M(i) |] ==> f \<in> ord_iso(A, r, i, Memrel(i))"
@@ -390,15 +390,15 @@
apply (blast intro: well_ord_ord_iso well_ord_Memrel)
done
-subsection {*Kunen's theorem 5.4, page 127*}
+subsection \<open>Kunen's theorem 5.4, page 127\<close>
-text{*(a) The notion of Wellordering is absolute*}
+text\<open>(a) The notion of Wellordering is absolute\<close>
theorem (in M_ordertype) well_ord_abs [simp]:
"[| M(A); M(r) |] ==> wellordered(M,A,r) \<longleftrightarrow> well_ord(A,r)"
by (blast intro: well_ord_imp_relativized relativized_imp_well_ord)
-text{*(b) Order types are absolute*}
+text\<open>(b) Order types are absolute\<close>
theorem (in M_ordertype)
"[| wellordered(M,A,r); f \<in> ord_iso(A, r, i, Memrel(i));
M(A); M(r); M(f); M(i); Ord(i) |] ==> i = ordertype(A,r)"
@@ -406,11 +406,11 @@
Ord_iso_implies_eq ord_iso_sym ord_iso_trans)
-subsection{*Ordinal Arithmetic: Two Examples of Recursion*}
+subsection\<open>Ordinal Arithmetic: Two Examples of Recursion\<close>
-text{*Note: the remainder of this theory is not needed elsewhere.*}
+text\<open>Note: the remainder of this theory is not needed elsewhere.\<close>
-subsubsection{*Ordinal Addition*}
+subsubsection\<open>Ordinal Addition\<close>
(*FIXME: update to use new techniques!!*)
(*This expresses ordinal addition in the language of ZF. It also
@@ -478,7 +478,7 @@
-text{*@{text is_oadd_fun}: Relating the pure "language of set theory" to Isabelle/ZF*}
+text\<open>@{text is_oadd_fun}: Relating the pure "language of set theory" to Isabelle/ZF\<close>
lemma (in M_ord_arith) is_oadd_fun_iff:
"[| a\<le>j; M(i); M(j); M(a); M(f) |]
==> is_oadd_fun(M,i,j,a,f) \<longleftrightarrow>
@@ -563,7 +563,7 @@
done
-subsubsection{*Ordinal Multiplication*}
+subsubsection\<open>Ordinal Multiplication\<close>
lemma omult_eqns_unique:
"[| omult_eqns(i,x,g,z); omult_eqns(i,x,g,z') |] ==> z=z'"
@@ -602,7 +602,7 @@
"[| M(i); M(x); M(g); function(g) |]
==> M(THE z. omult_eqns(i, x, g, z))"
apply (case_tac "Ord(x)")
- prefer 2 apply (simp add: omult_eqns_Not) --{*trivial, non-Ord case*}
+ prefer 2 apply (simp add: omult_eqns_Not) --\<open>trivial, non-Ord case\<close>
apply (erule Ord_cases)
apply (simp add: omult_eqns_0)
apply (simp add: omult_eqns_succ apply_closed oadd_closed)
@@ -669,11 +669,11 @@
-subsection {*Absoluteness of Well-Founded Relations*}
+subsection \<open>Absoluteness of Well-Founded Relations\<close>
-text{*Relativized to @{term M}: Every well-founded relation is a subset of some
+text\<open>Relativized to @{term M}: Every well-founded relation is a subset of some
inverse image of an ordinal. Key step is the construction (in @{term M}) of a
-rank function.*}
+rank function.\<close>
locale M_wfrank = M_trancl +
assumes wfrank_separation:
@@ -698,8 +698,8 @@
ordinal(M,rangef)))"
-text{*Proving that the relativized instances of Separation or Replacement
-agree with the "real" ones.*}
+text\<open>Proving that the relativized instances of Separation or Replacement
+agree with the "real" ones.\<close>
lemma (in M_wfrank) wfrank_separation':
"M(r) ==>
@@ -726,8 +726,8 @@
apply (simp add: relation2_def is_recfun_abs [of "%x. range"])
done
-text{*This function, defined using replacement, is a rank function for
-well-founded relations within the class M.*}
+text\<open>This function, defined using replacement, is a rank function for
+well-founded relations within the class M.\<close>
definition
wellfoundedrank :: "[i=>o,i,i] => i" where
"wellfoundedrank(M,r,A) ==
@@ -766,16 +766,16 @@
apply (rule wellfounded_induct, assumption, erule (1) transM)
apply simp
apply (blast intro: Ord_wfrank_separation', clarify)
-txt{*The reasoning in both cases is that we get @{term y} such that
+txt\<open>The reasoning in both cases is that we get @{term y} such that
@{term "\<langle>y, x\<rangle> \<in> r^+"}. We find that
- @{term "f`y = restrict(f, r^+ -`` {y})"}. *}
+ @{term "f`y = restrict(f, r^+ -`` {y})"}.\<close>
apply (rule OrdI [OF _ Ord_is_Transset])
- txt{*An ordinal is a transitive set...*}
+ txt\<open>An ordinal is a transitive set...\<close>
apply (simp add: Transset_def)
apply clarify
apply (frule apply_recfun2, assumption)
apply (force simp add: restrict_iff)
-txt{*...of ordinals. This second case requires the induction hyp.*}
+txt\<open>...of ordinals. This second case requires the induction hyp.\<close>
apply clarify
apply (rename_tac i y)
apply (frule apply_recfun2, assumption)
@@ -799,9 +799,9 @@
apply (simp add: wellfoundedrank_def)
apply (rule OrdI [OF _ Ord_is_Transset])
prefer 2
- txt{*by our previous result the range consists of ordinals.*}
+ txt\<open>by our previous result the range consists of ordinals.\<close>
apply (blast intro: Ord_wfrank_range)
-txt{*We still must show that the range is a transitive set.*}
+txt\<open>We still must show that the range is a transitive set.\<close>
apply (simp add: Transset_def, clarify, simp)
apply (rename_tac x i f u)
apply (frule is_recfun_imp_in_r, assumption)
@@ -814,7 +814,7 @@
apply (blast intro: is_recfun_restrict trans_trancl dest: apply_recfun2)
apply simp
apply blast
-txt{*Unicity requirement of Replacement*}
+txt\<open>Unicity requirement of Replacement\<close>
apply clarify
apply (frule apply_recfun2, assumption)
apply (simp add: trans_trancl is_recfun_cut)
@@ -824,7 +824,7 @@
"[| wellfounded(M,r); M(r); M(A)|]
==> function(wellfoundedrank(M,r,A))"
apply (simp add: wellfoundedrank_def function_def, clarify)
-txt{*Uniqueness: repeated below!*}
+txt\<open>Uniqueness: repeated below!\<close>
apply (drule is_recfun_functional, assumption)
apply (blast intro: wellfounded_trancl)
apply (simp_all add: trancl_subset_times trans_trancl)
@@ -841,7 +841,7 @@
apply (rule_tac x=x in ReplaceI)
apply simp
apply (rule_tac x=f in rexI, blast, simp_all)
-txt{*Uniqueness (for Replacement): repeated above!*}
+txt\<open>Uniqueness (for Replacement): repeated above!\<close>
apply clarify
apply (drule is_recfun_functional, assumption)
apply (blast intro: wellfounded_trancl)
@@ -875,7 +875,7 @@
apply (rule_tac x="range(f)" in rexI)
apply blast
apply simp_all
-txt{*Unicity requirement of Replacement*}
+txt\<open>Unicity requirement of Replacement\<close>
apply clarify
apply (drule is_recfun_functional, assumption)
apply (blast intro: wellfounded_trancl)
@@ -897,12 +897,12 @@
apply (frule is_recfun_restrict [of concl: "r^+" a])
apply (rule trans_trancl, assumption)
apply (simp_all add: r_into_trancl trancl_subset_times)
-txt{*Still the same goal, but with new @{text is_recfun} assumptions.*}
+txt\<open>Still the same goal, but with new @{text is_recfun} assumptions.\<close>
apply (simp add: wellfoundedrank_eq)
apply (frule_tac a=a in wellfoundedrank_eq, assumption+)
apply (simp_all add: transM [of a])
-txt{*We have used equations for wellfoundedrank and now must use some
- for @{text is_recfun}. *}
+txt\<open>We have used equations for wellfoundedrank and now must use some
+ for @{text is_recfun}.\<close>
apply (rule_tac a=a in rangeI)
apply (simp add: is_recfun_type [THEN apply_iff] vimage_singleton_iff
r_into_trancl apply_recfun r_into_trancl)
--- a/src/ZF/Constructible/Rank_Separation.thy Thu Jul 23 14:20:51 2015 +0200
+++ b/src/ZF/Constructible/Rank_Separation.thy Thu Jul 23 14:25:05 2015 +0200
@@ -2,19 +2,19 @@
Author: Lawrence C Paulson, Cambridge University Computer Laboratory
*)
-section {*Separation for Facts About Order Types, Rank Functions and
- Well-Founded Relations*}
+section \<open>Separation for Facts About Order Types, Rank Functions and
+ Well-Founded Relations\<close>
theory Rank_Separation imports Rank Rec_Separation begin
-text{*This theory proves all instances needed for locales
+text\<open>This theory proves all instances needed for locales
@{text "M_ordertype"} and @{text "M_wfrank"}. But the material is not
- needed for proving the relative consistency of AC.*}
+ needed for proving the relative consistency of AC.\<close>
-subsection{*The Locale @{text "M_ordertype"}*}
+subsection\<open>The Locale @{text "M_ordertype"}\<close>
-subsubsection{*Separation for Order-Isomorphisms*}
<