# HG changeset patch # User paulson # Date 1138012610 -3600 # Node ID 91a328803c6ac2bbfca0e19c3268791977a8813f # Parent 31c2af8b0c6051578b50fc4357b21361867db066 fixed the <<= notation diff -r 31c2af8b0c60 -r 91a328803c6a src/HOL/ex/Tarski.thy --- a/src/HOL/ex/Tarski.thy Mon Jan 23 11:36:05 2006 +0100 +++ b/src/HOL/ex/Tarski.thy Mon Jan 23 11:36:50 2006 +0100 @@ -85,10 +85,10 @@ (| pset = S, order = induced S (order cl) |): CompleteLattice }" syntax - "@SL" :: "['a set, 'a potype] => bool" ("_ <\ _" [51,50]50) + "@SL" :: "['a set, 'a potype] => bool" ("_ <<= _" [51,50]50) translations - "S <\ cl" == "S : sublattice `` {cl}" + "S <<= cl" == "S : sublattice `` {cl}" constdefs dual :: "'a potype => 'a potype" @@ -303,12 +303,12 @@ subsection {* sublattice *} lemma (in PO) sublattice_imp_CL: - "S <\ cl ==> (| pset = S, order = induced S r |) \ CompleteLattice" + "S <<= cl ==> (| pset = S, order = induced S r |) \ CompleteLattice" by (simp add: sublattice_def CompleteLattice_def A_def r_def) lemma (in CL) sublatticeI: "[| S \ A; (| pset = S, order = induced S r |) \ CompleteLattice |] - ==> S <\ cl" + ==> S <<= cl" by (simp add: sublattice_def A_def r_def) @@ -659,7 +659,7 @@ lemma (in CLF) interval_is_sublattice: "[| a \ A; b \ A; interval r a b \ {} |] - ==> interval r a b <\ cl" + ==> interval r a b <<= cl" apply (rule sublatticeI) apply (simp add: interval_subset) apply (rule CompleteLatticeI) @@ -737,7 +737,7 @@ done lemma (in Tarski) lubY_in_A: "lub Y cl \ A" -by (simp add: Y_subset_A [THEN lub_in_lattice]) + by (rule Y_subset_A [THEN lub_in_lattice]) lemma (in Tarski) lubY_le_flubY: "(lub Y cl, f (lub Y cl)) \ r" apply (rule lub_least)