# HG changeset patch # User huffman # Date 1334754547 -7200 # Node ID 1f0ec5b8135a21911a63aa8e959151234b2ff589 # Parent b06be48923a4a079b54d85c9d898ad326eee3a94 add lemma Quotient_abs_induct diff -r b06be48923a4 -r 1f0ec5b8135a src/HOL/Lifting.thy --- a/src/HOL/Lifting.thy Wed Apr 18 14:59:04 2012 +0200 +++ b/src/HOL/Lifting.thy Wed Apr 18 15:09:07 2012 +0200 @@ -296,6 +296,10 @@ lemma Quotient_rel_eq_transfer: "(T ===> T ===> op =) R (op =)" using 1 unfolding Quotient_alt_def fun_rel_def by simp +lemma Quotient_abs_induct: + assumes "\y. R y y \ P (Abs y)" shows "P x" + using 1 assms unfolding Quotient_def by metis + end text {* Generating transfer rules for total quotients. *}