src/LCF/ex/Ex1.ML
author wenzelm
Wed, 02 Nov 2005 14:46:49 +0100
changeset 18056 397b39b06ec8
parent 17878 5b9efe4d6b47
permissions -rw-r--r--
Sign.const_monomorphic;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
4905
be73ddff6c5a proper thy files;
wenzelm
parents:
diff changeset
     1
17878
wenzelm
parents: 17248
diff changeset
     2
(* $Id$ *)
4905
be73ddff6c5a proper thy files;
wenzelm
parents:
diff changeset
     3
be73ddff6c5a proper thy files;
wenzelm
parents:
diff changeset
     4
Addsimps [P_strict,K];
be73ddff6c5a proper thy files;
wenzelm
parents:
diff changeset
     5
17248
81bf91654e73 converted to Isar theory format;
wenzelm
parents: 5061
diff changeset
     6
val H_unfold = prove_goal (the_context ()) "H = K(H)"
4905
be73ddff6c5a proper thy files;
wenzelm
parents:
diff changeset
     7
 (fn _ => [stac H 1, rtac (FIX_eq RS sym) 1]);
be73ddff6c5a proper thy files;
wenzelm
parents:
diff changeset
     8
be73ddff6c5a proper thy files;
wenzelm
parents:
diff changeset
     9
bind_thm ("H_unfold", H_unfold);
be73ddff6c5a proper thy files;
wenzelm
parents:
diff changeset
    10
be73ddff6c5a proper thy files;
wenzelm
parents:
diff changeset
    11
17248
81bf91654e73 converted to Isar theory format;
wenzelm
parents: 5061
diff changeset
    12
val H_strict = prove_goal (the_context ()) "H(UU)=UU"
4905
be73ddff6c5a proper thy files;
wenzelm
parents:
diff changeset
    13
 (fn _ => [stac H_unfold 1, Simp_tac 1]);
be73ddff6c5a proper thy files;
wenzelm
parents:
diff changeset
    14
be73ddff6c5a proper thy files;
wenzelm
parents:
diff changeset
    15
bind_thm ("H_strict", H_strict);
be73ddff6c5a proper thy files;
wenzelm
parents:
diff changeset
    16
Addsimps [H_strict];
be73ddff6c5a proper thy files;
wenzelm
parents:
diff changeset
    17
5061
f947332d5465 isatool fixgoal;
wenzelm
parents: 4905
diff changeset
    18
Goal "ALL x. H(FIX(K,x)) = FIX(K,x)";
4905
be73ddff6c5a proper thy files;
wenzelm
parents:
diff changeset
    19
by (induct_tac "K" 1);
be73ddff6c5a proper thy files;
wenzelm
parents:
diff changeset
    20
by (Simp_tac 1);
be73ddff6c5a proper thy files;
wenzelm
parents:
diff changeset
    21
by (simp_tac (simpset() setloop (split_tac [COND_cases_iff])) 1);
be73ddff6c5a proper thy files;
wenzelm
parents:
diff changeset
    22
by (strip_tac 1);
be73ddff6c5a proper thy files;
wenzelm
parents:
diff changeset
    23
by (stac H_unfold 1);
be73ddff6c5a proper thy files;
wenzelm
parents:
diff changeset
    24
by (Asm_simp_tac 1);
be73ddff6c5a proper thy files;
wenzelm
parents:
diff changeset
    25
qed "H_idemp_lemma";
be73ddff6c5a proper thy files;
wenzelm
parents:
diff changeset
    26
be73ddff6c5a proper thy files;
wenzelm
parents:
diff changeset
    27
val H_idemp = rewrite_rule [mk_meta_eq (H RS sym)] H_idemp_lemma;
be73ddff6c5a proper thy files;
wenzelm
parents:
diff changeset
    28
bind_thm ("H_idemp", H_idemp);