diff -r 143f58bb34f9 -r 0909deb8059b src/HOL/Nominal/Examples/Support.thy --- a/src/HOL/Nominal/Examples/Support.thy Thu May 26 16:57:14 2016 +0200 +++ b/src/HOL/Nominal/Examples/Support.thy Thu May 26 17:51:22 2016 +0200 @@ -2,7 +2,7 @@ imports "../Nominal" begin -text {* +text \ An example showing that in general x\(A \ B) does not imply x\A and x\B @@ -11,31 +11,31 @@ the set of odd atoms. Then A \ B, that is the set of all atoms, has empty support. The sets A, respectively B, however have the set of all atoms as their support. -*} +\ atom_decl atom -text {* The set of even atoms. *} +text \The set of even atoms.\ abbreviation EVEN :: "atom set" where "EVEN \ {atom n | n. \i. n=2*i}" -text {* The set of odd atoms: *} +text \The set of odd atoms:\ abbreviation ODD :: "atom set" where "ODD \ {atom n | n. \i. n=2*i+1}" -text {* An atom is either even or odd. *} +text \An atom is either even or odd.\ lemma even_or_odd: fixes n :: nat shows "\i. (n = 2*i) \ (n=2*i+1)" by (induct n) (presburger)+ -text {* +text \ The union of even and odd atoms is the set of all atoms. - (Unfortunately I do not know a simpler proof of this fact.) *} + (Unfortunately I do not know a simpler proof of this fact.)\ lemma EVEN_union_ODD: shows "EVEN \ ODD = UNIV" using even_or_odd @@ -49,15 +49,15 @@ finally show "EVEN \ ODD = UNIV" by simp qed -text {* The sets of even and odd atoms are disjunct. *} +text \The sets of even and odd atoms are disjunct.\ lemma EVEN_intersect_ODD: shows "EVEN \ ODD = {}" using even_or_odd by (auto) (presburger) -text {* +text \ The preceeding two lemmas help us to prove - the following two useful equalities: *} + the following two useful equalities:\ lemma UNIV_subtract: shows "UNIV - EVEN = ODD" @@ -65,7 +65,7 @@ using EVEN_union_ODD EVEN_intersect_ODD by (blast)+ -text {* The sets EVEN and ODD are infinite. *} +text \The sets EVEN and ODD are infinite.\ lemma EVEN_ODD_infinite: shows "infinite EVEN" and "infinite ODD" @@ -80,10 +80,10 @@ then show "\f::nat\atom. inj f \ range f \ ODD" by (rule_tac exI) qed -text {* +text \ A general fact about a set S of atoms that is both infinite and coinfinite. Then S has all atoms as its support. Steve Zdancewic - helped with proving this fact. *} + helped with proving this fact.\ lemma supp_infinite_coinfinite: fixes S::"atom set" @@ -114,16 +114,16 @@ then show "(supp S) = (UNIV::atom set)" by auto qed -text {* As a corollary we get that EVEN and ODD have infinite support. *} +text \As a corollary we get that EVEN and ODD have infinite support.\ lemma EVEN_ODD_supp: shows "supp EVEN = (UNIV::atom set)" and "supp ODD = (UNIV::atom set)" using supp_infinite_coinfinite UNIV_subtract EVEN_ODD_infinite by simp_all -text {* +text \ The set of all atoms has empty support, since any swappings leaves - this set unchanged. *} + this set unchanged.\ lemma UNIV_supp: shows "supp (UNIV::atom set) = ({}::atom set)" @@ -133,7 +133,7 @@ then show "supp (UNIV::atom set) = ({}::atom set)" by (simp add: supp_def) qed -text {* Putting everything together. *} +text \Putting everything together.\ theorem EVEN_ODD_freshness: fixes x::"atom" shows "x\(EVEN \ ODD)" @@ -141,6 +141,6 @@ and "\x\ODD" by (auto simp only: fresh_def EVEN_union_ODD EVEN_ODD_supp UNIV_supp) -text {* Moral: support is a sublte notion. *} +text \Moral: support is a sublte notion.\ end