6008
|
1 |
(* Title: HOL/IOA/example/Correctness.thy
|
|
2 |
ID: $Id$
|
12218
|
3 |
Author: Olaf Müller
|
6008
|
4 |
*)
|
|
5 |
|
17244
|
6 |
header {* Correctness Proof *}
|
6008
|
7 |
|
17244
|
8 |
theory Correctness
|
|
9 |
imports SimCorrectness Spec Impl
|
|
10 |
begin
|
6008
|
11 |
|
17244
|
12 |
defaultsort type
|
6008
|
13 |
|
17244
|
14 |
constdefs
|
|
15 |
sim_relation :: "((nat * bool) * (nat set * bool)) set"
|
|
16 |
"sim_relation == {qua. let c = fst qua; a = snd qua ;
|
6008
|
17 |
k = fst c; b = snd c;
|
|
18 |
used = fst a; c = snd a
|
|
19 |
in
|
|
20 |
(! l:used. l < k) & b=c }"
|
|
21 |
|
19740
|
22 |
declare split_paired_All [simp]
|
|
23 |
declare split_paired_Ex [simp del]
|
|
24 |
|
|
25 |
|
|
26 |
(* Idea: instead of impl_con_lemma do not rewrite impl_ioa, but derive
|
|
27 |
simple lemmas asig_of impl_ioa = impl_sig, trans_of impl_ioa = impl_trans
|
|
28 |
Idea: ?ex. move .. should be generally replaced by a step via a subst tac if desired,
|
|
29 |
as this can be done globally *)
|
|
30 |
|
|
31 |
lemma issimulation:
|
|
32 |
"is_simulation sim_relation impl_ioa spec_ioa"
|
|
33 |
apply (simp (no_asm) add: is_simulation_def)
|
|
34 |
apply (rule conjI)
|
|
35 |
txt {* start states *}
|
|
36 |
apply (tactic "SELECT_GOAL (safe_tac set_cs) 1")
|
|
37 |
apply (rule_tac x = " ({},False) " in exI)
|
|
38 |
apply (simp add: sim_relation_def starts_of_def Spec.ioa_def Impl.ioa_def)
|
|
39 |
txt {* main-part *}
|
|
40 |
apply (rule allI)+
|
|
41 |
apply (rule imp_conj_lemma)
|
|
42 |
apply (rename_tac k b used c k' b' a)
|
|
43 |
apply (induct_tac "a")
|
|
44 |
apply (simp_all (no_asm) add: sim_relation_def Impl.ioa_def Impl.trans_def trans_of_def)
|
|
45 |
apply (tactic "safe_tac set_cs")
|
|
46 |
txt {* NEW *}
|
|
47 |
apply (rule_tac x = "(used,True)" in exI)
|
|
48 |
apply simp
|
|
49 |
apply (rule transition_is_ex)
|
|
50 |
apply (simp (no_asm) add: Spec.ioa_def Spec.trans_def trans_of_def)
|
|
51 |
txt {* LOC *}
|
|
52 |
apply (rule_tac x = " (used Un {k},False) " in exI)
|
|
53 |
apply (simp add: less_SucI)
|
|
54 |
apply (rule transition_is_ex)
|
|
55 |
apply (simp (no_asm) add: Spec.ioa_def Spec.trans_def trans_of_def)
|
|
56 |
apply fast
|
|
57 |
txt {* FREE *}
|
|
58 |
apply (rule_tac x = " (used - {nat},c) " in exI)
|
|
59 |
apply simp
|
|
60 |
apply (rule transition_is_ex)
|
|
61 |
apply (simp (no_asm) add: Spec.ioa_def Spec.trans_def trans_of_def)
|
|
62 |
done
|
|
63 |
|
|
64 |
|
|
65 |
lemma implementation:
|
|
66 |
"impl_ioa =<| spec_ioa"
|
|
67 |
apply (unfold ioa_implements_def)
|
|
68 |
apply (rule conjI)
|
|
69 |
apply (simp (no_asm) add: Impl.sig_def Spec.sig_def Impl.ioa_def Spec.ioa_def
|
|
70 |
asig_outputs_def asig_of_def asig_inputs_def)
|
|
71 |
apply (rule trace_inclusion_for_simulations)
|
|
72 |
apply (simp (no_asm) add: Impl.sig_def Spec.sig_def Impl.ioa_def Spec.ioa_def
|
|
73 |
externals_def asig_outputs_def asig_of_def asig_inputs_def)
|
|
74 |
apply (rule issimulation)
|
|
75 |
done
|
17244
|
76 |
|
|
77 |
end
|