src/HOL/UNITY/TimerArray.ML
author paulson
Tue, 21 Sep 1999 11:11:09 +0200
changeset 7547 a72a551b6d79
parent 7537 875754b599df
child 8041 e3237d8c18d6
permissions -rw-r--r--
new proof of drop_prog_correct for new definition of project_act
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
7513
879ae27f5e6f new example HOL/UNITY/TimerArray
paulson
parents:
diff changeset
     1
(*  Title:      HOL/UNITY/TimerArray.thy
879ae27f5e6f new example HOL/UNITY/TimerArray
paulson
parents:
diff changeset
     2
    ID:         $Id$
879ae27f5e6f new example HOL/UNITY/TimerArray
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
879ae27f5e6f new example HOL/UNITY/TimerArray
paulson
parents:
diff changeset
     4
    Copyright   1998  University of Cambridge
879ae27f5e6f new example HOL/UNITY/TimerArray
paulson
parents:
diff changeset
     5
879ae27f5e6f new example HOL/UNITY/TimerArray
paulson
parents:
diff changeset
     6
A trivial example of reasoning about an array of processes
879ae27f5e6f new example HOL/UNITY/TimerArray
paulson
parents:
diff changeset
     7
*)
879ae27f5e6f new example HOL/UNITY/TimerArray
paulson
parents:
diff changeset
     8
879ae27f5e6f new example HOL/UNITY/TimerArray
paulson
parents:
diff changeset
     9
Addsimps [Timer_def RS def_prg_Init];
879ae27f5e6f new example HOL/UNITY/TimerArray
paulson
parents:
diff changeset
    10
program_defs_ref := [Timer_def];
879ae27f5e6f new example HOL/UNITY/TimerArray
paulson
parents:
diff changeset
    11
7537
875754b599df working snapshot
paulson
parents: 7513
diff changeset
    12
Addsimps [decr_def];
7513
879ae27f5e6f new example HOL/UNITY/TimerArray
paulson
parents:
diff changeset
    13
879ae27f5e6f new example HOL/UNITY/TimerArray
paulson
parents:
diff changeset
    14
(*Demonstrates induction, but not used in the following proof*)
879ae27f5e6f new example HOL/UNITY/TimerArray
paulson
parents:
diff changeset
    15
Goal "Timer : UNIV leadsTo {0}";
879ae27f5e6f new example HOL/UNITY/TimerArray
paulson
parents:
diff changeset
    16
by (res_inst_tac [("f", "id")] (allI RS lessThan_induct) 1);
879ae27f5e6f new example HOL/UNITY/TimerArray
paulson
parents:
diff changeset
    17
by (Simp_tac 1);
879ae27f5e6f new example HOL/UNITY/TimerArray
paulson
parents:
diff changeset
    18
by (exhaust_tac "m" 1);
879ae27f5e6f new example HOL/UNITY/TimerArray
paulson
parents:
diff changeset
    19
by (blast_tac (claset() addSIs [subset_imp_leadsTo]) 1);
7537
875754b599df working snapshot
paulson
parents: 7513
diff changeset
    20
by (ensures_tac "decr" 1);
7513
879ae27f5e6f new example HOL/UNITY/TimerArray
paulson
parents:
diff changeset
    21
qed "Timer_leadsTo_zero";
879ae27f5e6f new example HOL/UNITY/TimerArray
paulson
parents:
diff changeset
    22
879ae27f5e6f new example HOL/UNITY/TimerArray
paulson
parents:
diff changeset
    23
Goal "finite I ==> (plam i: I. Timer) : UNIV leadsTo {s. ALL i:I. s i = 0}";
879ae27f5e6f new example HOL/UNITY/TimerArray
paulson
parents:
diff changeset
    24
by (eres_inst_tac [("A'1", "%i. lift_set i {0}")]
879ae27f5e6f new example HOL/UNITY/TimerArray
paulson
parents:
diff changeset
    25
    (finite_stable_completion RS leadsTo_weaken) 1);
879ae27f5e6f new example HOL/UNITY/TimerArray
paulson
parents:
diff changeset
    26
by Auto_tac;
7537
875754b599df working snapshot
paulson
parents: 7513
diff changeset
    27
by (constrains_tac 2);
7513
879ae27f5e6f new example HOL/UNITY/TimerArray
paulson
parents:
diff changeset
    28
by (res_inst_tac [("f", "sub i")] (allI RS lessThan_induct) 1);
879ae27f5e6f new example HOL/UNITY/TimerArray
paulson
parents:
diff changeset
    29
by (exhaust_tac "m" 1);
879ae27f5e6f new example HOL/UNITY/TimerArray
paulson
parents:
diff changeset
    30
by (auto_tac (claset() addIs [subset_imp_leadsTo], 
879ae27f5e6f new example HOL/UNITY/TimerArray
paulson
parents:
diff changeset
    31
	      simpset() addsimps [insert_absorb, lessThan_Suc RS sym]));
879ae27f5e6f new example HOL/UNITY/TimerArray
paulson
parents:
diff changeset
    32
by (rename_tac "n" 1);
7537
875754b599df working snapshot
paulson
parents: 7513
diff changeset
    33
by (rtac PLam_leadsTo_Basis 1);
7513
879ae27f5e6f new example HOL/UNITY/TimerArray
paulson
parents:
diff changeset
    34
by (auto_tac (claset(), simpset() addsimps [lessThan_Suc RS sym]));
879ae27f5e6f new example HOL/UNITY/TimerArray
paulson
parents:
diff changeset
    35
by (constrains_tac 1);
7537
875754b599df working snapshot
paulson
parents: 7513
diff changeset
    36
by (res_inst_tac [("act", "decr")] transient_mem 1);
7513
879ae27f5e6f new example HOL/UNITY/TimerArray
paulson
parents:
diff changeset
    37
by (auto_tac (claset(), simpset() addsimps [Timer_def]));
879ae27f5e6f new example HOL/UNITY/TimerArray
paulson
parents:
diff changeset
    38
qed "TimerArray_leadsTo_zero";