src/HOLCF/One.ML
author nipkow
Mon, 06 Aug 2001 13:43:24 +0200
changeset 11464 ddea204de5bc
parent 9248 e1dee89de037
child 12030 46d57d0290a2
permissions -rw-r--r--
turned translation for 1::nat into def. introduced 1' and replaced most occurrences of 1 by 1'.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2275
dbce3dce821a renamed is_flat to flat,
oheimb
parents: 1461
diff changeset
     1
(*  Title:      HOLCF/One.ML
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
     2
    ID:         $Id$
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2452
diff changeset
     3
    Author:     Oscar Slotosch
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2452
diff changeset
     4
    Copyright   1997 Technische Universitaet Muenchen
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
     5
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 5087
diff changeset
     6
The unit domain
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
(* ------------------------------------------------------------------------ *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    10
(* Exhaustion and Elimination for type one                                  *)
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
9248
e1dee89de037 massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents: 9245
diff changeset
    13
Goalw [ONE_def] "t=UU | t = ONE";
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 5087
diff changeset
    14
by (lift.induct_tac "t" 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 5087
diff changeset
    15
by (Simp_tac 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 5087
diff changeset
    16
by (Simp_tac 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 5087
diff changeset
    17
qed "Exh_one";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    18
9248
e1dee89de037 massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents: 9245
diff changeset
    19
val prems = Goal "[| p=UU ==> Q; p = ONE ==>Q|] ==>Q";
9245
428385c4bc50 removed most batch-style proofs
paulson
parents: 5087
diff changeset
    20
by (rtac (Exh_one RS disjE) 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 5087
diff changeset
    21
by (eresolve_tac prems 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 5087
diff changeset
    22
by (eresolve_tac prems 1);
428385c4bc50 removed most batch-style proofs
paulson
parents: 5087
diff changeset
    23
qed "oneE";
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    24
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2452
diff changeset
    25
(* ------------------------------------------------------------------------ *) 
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2452
diff changeset
    26
(* tactic for one-thms                                                      *)
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2452
diff changeset
    27
(* ------------------------------------------------------------------------ *)
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2452
diff changeset
    28
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2452
diff changeset
    29
fun prover t = prove_goalw thy [ONE_def] t
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2452
diff changeset
    30
 (fn prems =>
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2452
diff changeset
    31
        [
4098
71e05eb27fb6 isatool fixclasimp;
wenzelm
parents: 2640
diff changeset
    32
	(asm_simp_tac (simpset() addsimps [inst_lift_po]) 1)
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2452
diff changeset
    33
	]);
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2452
diff changeset
    34
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    35
(* ------------------------------------------------------------------------ *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    36
(* distinctness for type one : stored in a list                             *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    37
(* ------------------------------------------------------------------------ *)
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    38
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2452
diff changeset
    39
val dist_less_one = map prover ["~ONE << UU"];
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    40
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2452
diff changeset
    41
val dist_eq_one = map prover ["ONE~=UU","UU~=ONE"];
243
c22b85994e17 Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff changeset
    42
2640
ee4dfce170a0 Changes of HOLCF from Oscar Slotosch:
slotosch
parents: 2452
diff changeset
    43
Addsimps (dist_less_one@dist_eq_one);