diff -r 2c84860c19d2 -r 2ad8e66de294 Quot/Examples/LarryInt.thy --- a/Quot/Examples/LarryInt.thy Mon Feb 15 16:37:48 2010 +0100 +++ b/Quot/Examples/LarryInt.thy Mon Feb 15 16:51:30 2010 +0100 @@ -1,7 +1,7 @@ header{*The Integers as Equivalence Classes over Pairs of Natural Numbers*} -theory LarryInt +theory LarryInt imports Nat "../Quotient" "../Quotient_Product" begin @@ -373,12 +373,14 @@ "nat_raw" abbreviation - "less_raw x y \ (le_raw x y \ \(x = y))" + "less_raw x y \ (le_raw x y \ \(intrel x y))" lemma nat_le_eq_zle_raw: shows "less_raw (0, 0) w \ le_raw (0, 0) z \ (nat_raw w \ nat_raw z) = (le_raw w z)" -apply(auto) -sorry + apply (cases w) + apply (cases z) + apply (simp add: nat_raw_def le_raw_def) + by auto lemma [quot_respect]: shows "(intrel ===> op =) nat_raw nat_raw"