src/HOLCF/ex/Dnat.ML
author nipkow
Tue, 09 Jan 2001 15:36:30 +0100
changeset 10835 f4745d77e620
parent 9265 35aab1c9c08c
permissions -rw-r--r--
` -> $
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2570
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
     1
(*  Title:      HOLCF/Dnat.ML
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
     2
    ID:         $Id$
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
     3
    Author:     Franz Regensburger
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
     4
    Copyright   1993, 1995 Technische Universitaet Muenchen
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
     5
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
     6
*)
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
     7
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
     8
open Dnat;
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
     9
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    10
(* ------------------------------------------------------------------------- *)
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    11
(* expand fixed point properties                                             *)
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    12
(* ------------------------------------------------------------------------- *)
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    13
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    14
bind_thm ("iterator_def2", fix_prover2 Dnat.thy iterator_def 
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    15
	"iterator = (LAM n f x. case n of dzero => x | \
10835
nipkow
parents: 9265
diff changeset
    16
\                                     dsucc$m => f$(iterator$m$f$x))");
2570
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    17
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    18
(* ------------------------------------------------------------------------- *)
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    19
(* recursive  properties                                                     *)
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    20
(* ------------------------------------------------------------------------- *)
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    21
10835
nipkow
parents: 9265
diff changeset
    22
Goal "iterator$UU$f$x = UU";
9265
35aab1c9c08c removal of batch style, and tidying
paulson
parents: 4423
diff changeset
    23
by (stac iterator_def2 1);
35aab1c9c08c removal of batch style, and tidying
paulson
parents: 4423
diff changeset
    24
by (simp_tac (HOLCF_ss addsimps dnat.rews) 1);
35aab1c9c08c removal of batch style, and tidying
paulson
parents: 4423
diff changeset
    25
qed "iterator1";
2570
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    26
10835
nipkow
parents: 9265
diff changeset
    27
Goal "iterator$dzero$f$x = x";
9265
35aab1c9c08c removal of batch style, and tidying
paulson
parents: 4423
diff changeset
    28
by (stac iterator_def2 1);
35aab1c9c08c removal of batch style, and tidying
paulson
parents: 4423
diff changeset
    29
by (simp_tac (HOLCF_ss addsimps dnat.rews) 1);
35aab1c9c08c removal of batch style, and tidying
paulson
parents: 4423
diff changeset
    30
qed "iterator2";
2570
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    31
10835
nipkow
parents: 9265
diff changeset
    32
Goal "n~=UU ==> iterator$(dsucc$n)$f$x = f$(iterator$n$f$x)";
9265
35aab1c9c08c removal of batch style, and tidying
paulson
parents: 4423
diff changeset
    33
by (rtac trans 1);
35aab1c9c08c removal of batch style, and tidying
paulson
parents: 4423
diff changeset
    34
by (stac iterator_def2 1);
35aab1c9c08c removal of batch style, and tidying
paulson
parents: 4423
diff changeset
    35
by (asm_simp_tac (HOLCF_ss addsimps dnat.rews) 1);
35aab1c9c08c removal of batch style, and tidying
paulson
parents: 4423
diff changeset
    36
by (rtac refl 1);
35aab1c9c08c removal of batch style, and tidying
paulson
parents: 4423
diff changeset
    37
qed "iterator3";
2570
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    38
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    39
val iterator_rews = 
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    40
	[iterator1, iterator2, iterator3];
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    41
9265
35aab1c9c08c removal of batch style, and tidying
paulson
parents: 4423
diff changeset
    42
Goal "ALL x y::dnat. x<<y --> x=UU | x=y";
35aab1c9c08c removal of batch style, and tidying
paulson
parents: 4423
diff changeset
    43
by (rtac allI 1);
35aab1c9c08c removal of batch style, and tidying
paulson
parents: 4423
diff changeset
    44
by (res_inst_tac [("x","x")] dnat.ind 1);
35aab1c9c08c removal of batch style, and tidying
paulson
parents: 4423
diff changeset
    45
by (fast_tac HOL_cs 1);
35aab1c9c08c removal of batch style, and tidying
paulson
parents: 4423
diff changeset
    46
by (rtac allI 1);
35aab1c9c08c removal of batch style, and tidying
paulson
parents: 4423
diff changeset
    47
by (res_inst_tac [("x","y")] dnat.casedist 1);
35aab1c9c08c removal of batch style, and tidying
paulson
parents: 4423
diff changeset
    48
by (fast_tac (HOL_cs addSIs [UU_I]) 1);
35aab1c9c08c removal of batch style, and tidying
paulson
parents: 4423
diff changeset
    49
by (Asm_simp_tac 1);
35aab1c9c08c removal of batch style, and tidying
paulson
parents: 4423
diff changeset
    50
by (asm_simp_tac (simpset() addsimps dnat.dist_les) 1);
35aab1c9c08c removal of batch style, and tidying
paulson
parents: 4423
diff changeset
    51
by (rtac allI 1);
35aab1c9c08c removal of batch style, and tidying
paulson
parents: 4423
diff changeset
    52
by (res_inst_tac [("x","y")] dnat.casedist 1);
35aab1c9c08c removal of batch style, and tidying
paulson
parents: 4423
diff changeset
    53
by (fast_tac (HOL_cs addSIs [UU_I]) 1);
35aab1c9c08c removal of batch style, and tidying
paulson
parents: 4423
diff changeset
    54
by (asm_simp_tac (simpset() addsimps dnat.dist_les) 1);
35aab1c9c08c removal of batch style, and tidying
paulson
parents: 4423
diff changeset
    55
by (asm_simp_tac (simpset() addsimps dnat.rews) 1);
35aab1c9c08c removal of batch style, and tidying
paulson
parents: 4423
diff changeset
    56
by (strip_tac 1);
35aab1c9c08c removal of batch style, and tidying
paulson
parents: 4423
diff changeset
    57
by (subgoal_tac "d<<da" 1);
35aab1c9c08c removal of batch style, and tidying
paulson
parents: 4423
diff changeset
    58
by (etac allE 1);
35aab1c9c08c removal of batch style, and tidying
paulson
parents: 4423
diff changeset
    59
by (dtac mp 1);
35aab1c9c08c removal of batch style, and tidying
paulson
parents: 4423
diff changeset
    60
by (atac 1);
35aab1c9c08c removal of batch style, and tidying
paulson
parents: 4423
diff changeset
    61
by (etac disjE 1);
35aab1c9c08c removal of batch style, and tidying
paulson
parents: 4423
diff changeset
    62
by (contr_tac 1);
35aab1c9c08c removal of batch style, and tidying
paulson
parents: 4423
diff changeset
    63
by (Asm_simp_tac 1);
35aab1c9c08c removal of batch style, and tidying
paulson
parents: 4423
diff changeset
    64
by (resolve_tac  dnat.inverts 1);
35aab1c9c08c removal of batch style, and tidying
paulson
parents: 4423
diff changeset
    65
by (REPEAT (atac 1));
35aab1c9c08c removal of batch style, and tidying
paulson
parents: 4423
diff changeset
    66
qed "dnat_flat";
2570
24d7e8fb8261 added Classlib.* and Witness.*,
oheimb
parents:
diff changeset
    67