| author | wenzelm | 
| Wed, 27 Oct 1999 17:27:07 +0200 | |
| changeset 7954 | ea6b79f32cfd | 
| parent 7661 | 8c3190b173aa | 
| child 9245 | 428385c4bc50 | 
| permissions | -rw-r--r-- | 
| 2640 | 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 | 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 | 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 | |
| 7661 | 16 | qed_goal "Sel_injective_cprod" Prod.thy | 
| 2640 | 17 | "[|fst x = fst y; snd x = snd y|] ==> x = y" | 
| 18 | (fn prems => | |
| 1461 | 19 | [ | 
| 20 | (cut_facts_tac prems 1), | |
| 2640 | 21 | (subgoal_tac "(fst x,snd x)=(fst y,snd y)" 1), | 
| 22 | (rotate_tac ~1 1), | |
| 23 | (asm_full_simp_tac(HOL_ss addsimps[surjective_pairing RS sym])1), | |
| 7661 | 24 | (asm_simp_tac (simpset_of Prod.thy) 1) | 
| 1461 | 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: 
2640diff
changeset | 27 | qed_goalw "refl_less_cprod" Cprod1.thy [less_cprod_def] "(p::'a*'b) << p" | 
| 2640 | 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 | 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: 
2640diff
changeset | 31 | "[|(p1::'a * 'b) << p2;p2 << p1|] ==> p1=p2" | 
| 2640 | 32 | (fn prems => | 
| 1461 | 33 | [ | 
| 34 | (cut_facts_tac prems 1), | |
| 2640 | 35 | (rtac Sel_injective_cprod 1), | 
| 36 | (fast_tac (HOL_cs addIs [antisym_less]) 1), | |
| 37 | (fast_tac (HOL_cs addIs [antisym_less]) 1) | |
| 1461 | 38 | ]); | 
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 39 | |
| 2640 | 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: 
2640diff
changeset | 41 | "[|(p1::'a*'b) << p2;p2 << p3|] ==> p1 << p3" | 
| 2640 | 42 | (fn prems => | 
| 43 | [ | |
| 44 | (cut_facts_tac prems 1), | |
| 45 | (rtac conjI 1), | |
| 46 | (fast_tac (HOL_cs addIs [trans_less]) 1), | |
| 47 | (fast_tac (HOL_cs addIs [trans_less]) 1) | |
| 48 | ]); |