diff -r f0aeca99b5d9 -r 461ee3e49ad3 src/HOL/Algebra/IntRing.thy --- a/src/HOL/Algebra/IntRing.thy Thu Mar 26 19:24:21 2009 +0100 +++ b/src/HOL/Algebra/IntRing.thy Thu Mar 26 20:08:55 2009 +0100 @@ -96,7 +96,7 @@ interpretation needs to be done as early as possible --- that is, with as few assumptions as possible. *} -interpretation int!: monoid \ +interpretation int: monoid \ where "carrier \ = UNIV" and "mult \ x y = x * y" and "one \ = 1" @@ -104,7 +104,7 @@ proof - -- "Specification" show "monoid \" proof qed (auto simp: int_ring_def) - then interpret int!: monoid \ . + then interpret int: monoid \ . -- "Carrier" show "carrier \ = UNIV" by (simp add: int_ring_def) @@ -116,12 +116,12 @@ show "pow \ x n = x^n" by (induct n) (simp, simp add: int_ring_def)+ qed -interpretation int!: comm_monoid \ +interpretation int: comm_monoid \ where "finprod \ f A = (if finite A then setprod f A else undefined)" proof - -- "Specification" show "comm_monoid \" proof qed (auto simp: int_ring_def) - then interpret int!: comm_monoid \ . + then interpret int: comm_monoid \ . -- "Operations" { fix x y have "mult \ x y = x * y" by (simp add: int_ring_def) } @@ -139,14 +139,14 @@ qed qed -interpretation int!: abelian_monoid \ +interpretation int: abelian_monoid \ where "zero \ = 0" and "add \ x y = x + y" and "finsum \ f A = (if finite A then setsum f A else undefined)" proof - -- "Specification" show "abelian_monoid \" proof qed (auto simp: int_ring_def) - then interpret int!: abelian_monoid \ . + then interpret int: abelian_monoid \ . -- "Operations" { fix x y show "add \ x y = x + y" by (simp add: int_ring_def) } @@ -164,7 +164,7 @@ qed qed -interpretation int!: abelian_group \ +interpretation int: abelian_group \ where "a_inv \ x = - x" and "a_minus \ x y = x - y" proof - @@ -174,7 +174,7 @@ show "!!x. x \ carrier \ ==> EX y : carrier \. y \\<^bsub>\\<^esub> x = \\<^bsub>\\<^esub>" by (simp add: int_ring_def) arith qed (auto simp: int_ring_def) - then interpret int!: abelian_group \ . + then interpret int: abelian_group \ . -- "Operations" { fix x y have "add \ x y = x + y" by (simp add: int_ring_def) } @@ -187,7 +187,7 @@ show "a_minus \ x y = x - y" by (simp add: int.minus_eq add a_inv) qed -interpretation int!: "domain" \ +interpretation int: "domain" \ proof qed (auto simp: int_ring_def left_distrib right_distrib) @@ -203,7 +203,7 @@ "(True ==> PROP R) == PROP R" by simp_all -interpretation int! (* FIXME [unfolded UNIV] *) : +interpretation int (* FIXME [unfolded UNIV] *) : partial_order "(| carrier = UNIV::int set, eq = op =, le = op \ |)" where "carrier (| carrier = UNIV::int set, eq = op =, le = op \ |) = UNIV" and "le (| carrier = UNIV::int set, eq = op =, le = op \ |) x y = (x \ y)" @@ -219,7 +219,7 @@ by (simp add: lless_def) auto qed -interpretation int! (* FIXME [unfolded UNIV] *) : +interpretation int (* FIXME [unfolded UNIV] *) : lattice "(| carrier = UNIV::int set, eq = op =, le = op \ |)" where "join (| carrier = UNIV::int set, eq = op =, le = op \ |) x y = max x y" and "meet (| carrier = UNIV::int set, eq = op =, le = op \ |) x y = min x y" @@ -232,7 +232,7 @@ apply (simp add: greatest_def Lower_def) apply arith done - then interpret int!: lattice "?Z" . + then interpret int: lattice "?Z" . show "join ?Z x y = max x y" apply (rule int.joinI) apply (simp_all add: least_def Upper_def) @@ -245,7 +245,7 @@ done qed -interpretation int! (* [unfolded UNIV] *) : +interpretation int (* [unfolded UNIV] *) : total_order "(| carrier = UNIV::int set, eq = op =, le = op \ |)" proof qed clarsimp