# HG changeset patch # User Cezary Kaliszyk # Date 1259914726 -3600 # Node ID 28bb34eeedc5ce593f180772135a3e3225eb33ce # Parent 8dbc521ee97f0b6f2652e03d22424c24dd922dfe Changing = to \ in case if we want to use simp. diff -r 8dbc521ee97f -r 28bb34eeedc5 QuotList.thy --- a/QuotList.thy Fri Dec 04 09:10:31 2009 +0100 +++ b/QuotList.thy Fri Dec 04 09:18:46 2009 +0100 @@ -15,7 +15,7 @@ | "LIST_REL R (x#xs) (y#ys) = (R x y \ LIST_REL R xs ys)" lemma LIST_REL_EQ: - shows "LIST_REL (op =) = (op =)" + shows "LIST_REL (op =) \ (op =)" unfolding expand_fun_eq apply(rule allI)+ apply(induct_tac x xa rule: list_induct2') diff -r 8dbc521ee97f -r 28bb34eeedc5 QuotMain.thy --- a/QuotMain.thy Fri Dec 04 09:10:31 2009 +0100 +++ b/QuotMain.thy Fri Dec 04 09:18:46 2009 +0100 @@ -956,7 +956,7 @@ (* (op =) ===> (op =) \ (op =), needed in order to apply respectfulness theorems *) (* global simplification *) - (*NDT ctxt "H" (CHANGED o (asm_full_simp_tac ((Simplifier.context ctxt empty_ss) addsimps @{thms eq_reflection[OF FUN_REL_EQ] eq_reflection[OF LIST_REL_EQ]})))*) + (*NDT ctxt "H" (CHANGED o (asm_full_simp_tac ((Simplifier.context ctxt empty_ss) addsimps @{thms FUN_REL_EQ LIST_REL_EQ})))*) ]) *} diff -r 8dbc521ee97f -r 28bb34eeedc5 QuotScript.thy --- a/QuotScript.thy Fri Dec 04 09:10:31 2009 +0100 +++ b/QuotScript.thy Fri Dec 04 09:18:46 2009 +0100 @@ -136,7 +136,7 @@ "E1 ===> E2 \ FUN_REL E1 E2" lemma FUN_REL_EQ: - "(op =) ===> (op =) = (op =)" + "(op =) ===> (op =) \ (op =)" by (simp add: expand_fun_eq) lemma FUN_QUOTIENT: