src/HOLCF/Cprod1.ML
author wenzelm
Sun, 29 Nov 1998 13:14:45 +0100
changeset 5986 6ebbc9e7cc20
parent 3323 194ae2e0c193
child 7661 8c3190b173aa
permissions -rw-r--r--
added oct_char;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2033
diff changeset
     1
(*  Title:      HOLCF/Cprod1.ML
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
     2
    ID:         $Id$
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1267
diff changeset
     3
    Author:     Franz Regensburger
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
     4
    Copyright   1993  Technische Universitaet Muenchen
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
     5
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2033
diff changeset
     6
Lemmas for theory Cprod1.thy 
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
     7
*)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
     8
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
     9
open Cprod1;
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    10
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    11
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    12
(* ------------------------------------------------------------------------ *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    13
(* less_cprod is a partial order on 'a * 'b                                 *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    14
(* ------------------------------------------------------------------------ *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    15
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2033
diff changeset
    16
qed_goal "Sel_injective_cprod" Prod.thy 
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2033
diff changeset
    17
        "[|fst x = fst y; snd x = snd y|] ==> x = y"
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2033
diff changeset
    18
(fn prems =>
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1267
diff changeset
    19
        [
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1267
diff changeset
    20
        (cut_facts_tac prems 1),
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2033
diff changeset
    21
        (subgoal_tac "(fst x,snd x)=(fst y,snd y)" 1),
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2033
diff changeset
    22
        (rotate_tac ~1 1),
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2033
diff changeset
    23
        (asm_full_simp_tac(HOL_ss addsimps[surjective_pairing RS sym])1),
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2033
diff changeset
    24
        (Asm_simp_tac 1)
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1267
diff changeset
    25
        ]);
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    26
3323
194ae2e0c193 eliminated the constant less by the introduction of the axclass sq_ord
slotosch
parents: 2640
diff changeset
    27
qed_goalw "refl_less_cprod" Cprod1.thy [less_cprod_def] "(p::'a*'b) << p"
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2033
diff changeset
    28
 (fn prems => [Simp_tac 1]);
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    29
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2033
diff changeset
    30
qed_goalw "antisym_less_cprod" thy [less_cprod_def]
3323
194ae2e0c193 eliminated the constant less by the introduction of the axclass sq_ord
slotosch
parents: 2640
diff changeset
    31
        "[|(p1::'a * 'b) << p2;p2 << p1|] ==> p1=p2"
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2033
diff changeset
    32
(fn prems =>
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1267
diff changeset
    33
        [
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1267
diff changeset
    34
        (cut_facts_tac prems 1),
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2033
diff changeset
    35
        (rtac Sel_injective_cprod 1),
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2033
diff changeset
    36
        (fast_tac (HOL_cs addIs [antisym_less]) 1),
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2033
diff changeset
    37
        (fast_tac (HOL_cs addIs [antisym_less]) 1)
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1267
diff changeset
    38
        ]);
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    39
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2033
diff changeset
    40
qed_goalw "trans_less_cprod" thy [less_cprod_def]
3323
194ae2e0c193 eliminated the constant less by the introduction of the axclass sq_ord
slotosch
parents: 2640
diff changeset
    41
        "[|(p1::'a*'b) << p2;p2 << p3|] ==> p1 << p3"
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2033
diff changeset
    42
(fn prems =>
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2033
diff changeset
    43
        [
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2033
diff changeset
    44
        (cut_facts_tac prems 1),
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2033
diff changeset
    45
        (rtac conjI 1),
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2033
diff changeset
    46
        (fast_tac (HOL_cs addIs [trans_less]) 1),
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2033
diff changeset
    47
        (fast_tac (HOL_cs addIs [trans_less]) 1)
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2033
diff changeset
    48
        ]);