diff -r 143f58bb34f9 -r 0909deb8059b src/HOL/Nominal/Examples/Class1.thy --- a/src/HOL/Nominal/Examples/Class1.thy Thu May 26 16:57:14 2016 +0200 +++ b/src/HOL/Nominal/Examples/Class1.thy Thu May 26 17:51:22 2016 +0200 @@ -2,11 +2,11 @@ imports "../Nominal" begin -section {* Term-Calculus from Urban's PhD *} +section \Term-Calculus from Urban's PhD\ atom_decl name coname -text {* types *} +text \types\ nominal_datatype ty = PR "string" @@ -44,7 +44,7 @@ by (nominal_induct T rule: ty.strong_induct) (auto simp add: fresh_string) -text {* terms *} +text \terms\ nominal_datatype trm = Ax "name" "coname" @@ -60,15 +60,15 @@ | ImpR "\name\(\coname\trm)" "coname" ("ImpR (_).<_>._ _" [100,100,100,100] 100) | ImpL "\coname\trm" "\name\trm" "name" ("ImpL <_>._ (_)._ _" [100,100,100,100,100] 100) -text {* named terms *} +text \named terms\ nominal_datatype ntrm = Na "\name\trm" ("((_):_)" [100,100] 100) -text {* conamed terms *} +text \conamed terms\ nominal_datatype ctrm = Co "\coname\trm" ("(<_>:_)" [100,100] 100) -text {* renaming functions *} +text \renaming functions\ nominal_primrec (freshness_context: "(d::coname,e::coname)") crename :: "trm \ coname \ coname \ trm" ("_[_\c>_]" [100,100,100] 100) @@ -307,7 +307,7 @@ apply(simp_all add: trm.inject split: if_splits) done -text {* substitution functions *} +text \substitution functions\ lemma fresh_perm_coname: fixes c::"coname" @@ -2835,7 +2835,7 @@ lemmas subst_comm' = substn_crename_comm' substc_crename_comm' substn_nrename_comm' substc_nrename_comm' -text {* typing contexts *} +text \typing contexts\ type_synonym ctxtn = "(name\ty) list" type_synonym ctxtc = "(coname\ty) list" @@ -2868,7 +2868,7 @@ show "x\\" by (induct \) (auto simp add: fresh_list_nil fresh_list_cons fresh_prod fresh_atm fresh_ty) qed -text {* cut-reductions *} +text \cut-reductions\ declare abs_perm[eqvt] @@ -4692,7 +4692,7 @@ finally show ?thesis by simp qed -text {* axioms do not reduce *} +text \axioms do not reduce\ lemma ax_do_not_l_reduce: shows "Ax x a \\<^sub>l M \ False" @@ -5160,7 +5160,7 @@ apply(simp add: calc_atm perm_swap) done -text {* Transitive Closure*} +text \Transitive Closure\ abbreviation a_star_redu :: "trm \ trm \ bool" ("_ \\<^sub>a* _" [100,100] 100) @@ -5189,7 +5189,7 @@ using a2 a1 by (induct) (auto) -text {* congruence rules for \\\<^sub>a*\ *} +text \congruence rules for \\\<^sub>a*\\ lemma ax_do_not_a_star_reduce: shows "Ax x a \\<^sub>a* M \ M = Ax x a" @@ -5379,7 +5379,7 @@ apply(auto simp add: alpha trm.inject) done -text {* Substitution *} +text \Substitution\ lemma subst_not_fin1: shows "\fin(M{x:=.P}) x" @@ -5644,7 +5644,7 @@ apply(drule fic_elims, simp) done -text {* Reductions *} +text \Reductions\ lemma fin_l_reduce: assumes a: "fin M x" @@ -5782,7 +5782,7 @@ apply(auto simp add: fic_a_reduce) done -text {* substitution properties *} +text \substitution properties\ lemma subst_with_ax1: shows "M{x:=.Ax y a} \\<^sub>a* M[x\n>y]" @@ -6295,7 +6295,7 @@ using fs by simp qed -text {* substitution lemmas *} +text \substitution lemmas\ lemma not_Ax1: shows "\(b\M) \ M{b:=(y).Q} \ Ax x a"