src/HOLCF/IOA/meta_theory/LiveIOA.ML
author aspinall
Fri, 30 Sep 2005 18:18:34 +0200
changeset 17740 fc385ce6187d
parent 17233 41eee2e7b465
child 19360 f47412f922ab
permissions -rw-r--r--
Add icon for interface.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
     1
(*  Title:      HOLCF/IOA/meta_theory/LiveIOA.ML
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
     2
    ID:         $Id$
12218
wenzelm
parents: 6161
diff changeset
     3
    Author:     Olaf Müller
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
     4
*)   
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
     5
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
     6
Delsimps [split_paired_Ex];
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
     7
5068
fb28eaa07e01 isatool fixgoal;
wenzelm
parents: 4559
diff changeset
     8
Goalw [live_implements_def] 
13388
eff0ede61da1 quantify LC (conflict with const name of HOL);
wenzelm
parents: 12218
diff changeset
     9
"!!LC. [| live_implements (A,LA) (B,LB); live_implements (B,LB) (C,LC) |] \
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    10
\     ==> live_implements (A,LA) (C,LC)"; 
5132
24f992a25adc isatool expandshort;
wenzelm
parents: 5068
diff changeset
    11
by Auto_tac;
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    12
qed"live_implements_trans";
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    13
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    14
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    15
section "Correctness of live refmap";
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    16
	
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    17
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    18
(* ---------------------------------------------------------------- *)
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    19
(*                Correctness of live refmap                        *)
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    20
(* ---------------------------------------------------------------- *)
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    21
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    22
6161
paulson
parents: 6023
diff changeset
    23
Goal "[| inp(C)=inp(A); out(C)=out(A); \
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    24
\                  is_live_ref_map f (C,M) (A,L) |] \
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    25
\               ==> live_implements (C,M) (A,L)";
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    26
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    27
by (asm_full_simp_tac (simpset() addsimps [is_live_ref_map_def, live_implements_def,
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    28
livetraces_def,liveexecutions_def]) 1);
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    29
by (safe_tac set_cs);
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    30
by (res_inst_tac[("x","corresp_ex A f ex")] exI 1);
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    31
by (safe_tac set_cs);
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    32
  (* Traces coincide, Lemma 1 *)
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    33
  by (pair_tac "ex" 1);
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    34
  by (etac (lemma_1 RS spec RS mp) 1);
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    35
  by (simp_tac (simpset() addsimps [externals_def])1);
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    36
  by (SELECT_GOAL (auto_tac (claset(),simpset()))1);
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    37
  by (asm_full_simp_tac (simpset() addsimps [executions_def,reachable.reachable_0]) 1);
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    38
 
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    39
  (* corresp_ex is execution, Lemma 2 *)
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    40
  by (pair_tac "ex" 1);
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    41
  by (asm_full_simp_tac (simpset() addsimps [executions_def]) 1);
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    42
  (* start state *) 
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    43
  by (rtac conjI 1);
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    44
  by (asm_full_simp_tac (simpset() addsimps [is_ref_map_def,corresp_ex_def]) 1);
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    45
  (* is-execution-fragment *)
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    46
  by (etac (lemma_2 RS spec RS mp) 1);
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    47
  by (asm_full_simp_tac (simpset() addsimps [reachable.reachable_0]) 1);
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    48
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    49
 (* Liveness *)
5132
24f992a25adc isatool expandshort;
wenzelm
parents: 5068
diff changeset
    50
by Auto_tac;
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents:
diff changeset
    51
qed"live_implements";