# HG changeset patch # User haftmann # Date 1207575453 -7200 # Node ID 36a93808642cf59f869fcba3dcf22a59826a9c70 # Parent 522f45a8604e8f70be7741e79f28e370957a8304 instantiation replacing primitive instance plus overloaded defs diff -r 522f45a8604e -r 36a93808642c src/HOL/Bali/Decl.thy --- a/src/HOL/Bali/Decl.thy Mon Apr 07 15:37:32 2008 +0200 +++ b/src/HOL/Bali/Decl.thy Mon Apr 07 15:37:33 2008 +0200 @@ -6,7 +6,9 @@ *} (** order is significant, because of clash for "var" **) -theory Decl imports Term Table begin +theory Decl +imports Term Table +begin text {* improvements: @@ -50,21 +52,21 @@ Private < Package < Protected < Public *} -instance acc_modi:: ord .. +instantiation acc_modi :: linorder +begin -defs (overloaded) -less_acc_def: - "a < (b::acc_modi) - \ (case a of +definition + less_acc_def: "a < b + \ (case a of Private \ (b=Package \ b=Protected \ b=Public) | Package \ (b=Protected \ b=Public) | Protected \ (b=Public) | Public \ False)" -le_acc_def: - "a \ (b::acc_modi) \ (a = b) \ (a < b)" -instance acc_modi:: order -proof +definition + le_acc_def: "a \ (b::acc_modi) \ (a = b) \ (a < b)" + +instance proof fix x y z::acc_modi { show "x \ x" \\ -- reflexivity @@ -86,14 +88,12 @@ show "(x < y) = (x \ y \ x \ y)" by (auto simp add: le_acc_def less_acc_def split add: acc_modi.split) } -qed - -instance acc_modi:: linorder -proof fix x y:: acc_modi show "x \ y \ y \ x" by (auto simp add: less_acc_def le_acc_def split add: acc_modi.split) qed + +end lemma acc_modi_top [simp]: "Public \ a \ a = Public" by (auto simp add: le_acc_def less_acc_def split: acc_modi.splits) diff -r 522f45a8604e -r 36a93808642c src/HOL/Hyperreal/HyperDef.thy --- a/src/HOL/Hyperreal/HyperDef.thy Mon Apr 07 15:37:32 2008 +0200 +++ b/src/HOL/Hyperreal/HyperDef.thy Mon Apr 07 15:37:33 2008 +0200 @@ -43,11 +43,16 @@ subsection {* Real vector class instances *} -instance star :: (scaleR) scaleR .. +instantiation star :: (scaleR) scaleR +begin -defs (overloaded) +definition star_scaleR_def [transfer_unfold]: "scaleR r \ *f* (scaleR r)" +instance .. + +end + lemma Standard_scaleR [simp]: "x \ Standard \ scaleR r x \ Standard" by (simp add: star_scaleR_def)