src/HOLCF/Cprod1.ML
author wenzelm
Thu, 15 Feb 2001 17:18:54 +0100
changeset 11145 3e47692e3a3e
parent 11025 a70b796d9af8
child 11343 d5f1b482bfbf
permissions -rw-r--r--
eliminate get_def;

(*  Title:      HOLCF/Cprod1.ML
    ID:         $Id$
    Author:     Franz Regensburger
    Copyright   1993  Technische Universitaet Muenchen

Partial ordering for cartesian product of HOL theory Product_Type.thy
*)


(* ------------------------------------------------------------------------ *)
(* less_cprod is a partial order on 'a * 'b                                 *)
(* ------------------------------------------------------------------------ *)

(*###TO Product_Type_lemmas.ML *)
Goal "[|fst x = fst y; snd x = snd y|] ==> x = y";
by (subgoal_tac "(fst x,snd x)=(fst y,snd y)" 1);
by (rotate_tac ~1 1);
by (asm_full_simp_tac(HOL_ss addsimps[surjective_pairing RS sym])1);
by (asm_simp_tac (simpset_of (theory "Product_Type")) 1);
qed "Sel_injective_cprod";

Goalw [less_cprod_def] "(p::'a*'b) << p";
by (Simp_tac 1);
qed "refl_less_cprod";

Goalw [less_cprod_def] "[|(p1::'a * 'b) << p2;p2 << p1|] ==> p1=p2";
by (rtac Sel_injective_cprod 1);
by (fast_tac (HOL_cs addIs [antisym_less]) 1);
by (fast_tac (HOL_cs addIs [antisym_less]) 1);
qed "antisym_less_cprod";

Goalw [less_cprod_def]
        "[|(p1::'a*'b) << p2;p2 << p3|] ==> p1 << p3";
by (rtac conjI 1);
by (fast_tac (HOL_cs addIs [trans_less]) 1);
by (fast_tac (HOL_cs addIs [trans_less]) 1);
qed "trans_less_cprod";