src/HOL/IMP/Natural.ML
author paulson
Thu, 05 Sep 1996 18:42:48 +0200
changeset 1958 6f20ecbd87cb
parent 1730 1c7f793fc374
child 1973 8c94c9a5be10
permissions -rw-r--r--
Now uses thin_tac
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1700
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
     1
(*  Title:      HOL/IMP/Natural.ML
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
     2
    ID:         $Id$
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
     3
    Author:     Tobias Nipkow & Larry Paulson, TUM
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
     4
    Copyright   1996 TUM
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
     5
*)
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
     6
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
     7
open Natural;
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
     8
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
     9
val evalc_elim_cases = map (evalc.mk_cases com.simps)
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
    10
   ["<SKIP,s> -c-> t", "<x:=a,s> -c-> t", "<c1;c2, s> -c-> t",
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
    11
    "<IF b THEN c1 ELSE c2, s> -c-> t"];
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
    12
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
    13
val evalc_WHILE_case = evalc.mk_cases com.simps "<WHILE b DO c,s> -c-> t";
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
    14
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
    15
val evalc_cs = set_cs addSEs evalc_elim_cases
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
    16
                      addEs  [evalc_WHILE_case];
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
    17
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
    18
(* evaluation of com is deterministic *)
1730
1c7f793fc374 Updated for new form of induction rules
paulson
parents: 1700
diff changeset
    19
goal Natural.thy "!!c s t. <c,s> -c-> t ==> (!u. <c,s> -c-> u --> u=t)";
1c7f793fc374 Updated for new form of induction rules
paulson
parents: 1700
diff changeset
    20
by (etac evalc.induct 1);
1958
6f20ecbd87cb Now uses thin_tac
paulson
parents: 1730
diff changeset
    21
by (thin_tac "<?c,s2> -c-> s1" 7);
1700
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
    22
by (ALLGOALS (deepen_tac evalc_cs 4));
afd3b60660db Natural and Transition semantics.
nipkow
parents:
diff changeset
    23
qed_spec_mp "com_det";