src/HOL/HOLCF/IOA/RefMappings.thy
author wenzelm
Mon, 25 Jul 2016 21:50:04 +0200
changeset 63549 b0d31c7def86
parent 62390 842917225d56
permissions -rw-r--r--
more symbols;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
62008
cbedaddc9351 clarified directory structure;
wenzelm
parents: 62003
diff changeset
     1
(*  Title:      HOL/HOLCF/IOA/RefMappings.thy
40945
b8703f63bfb2 recoded latin1 as utf8;
wenzelm
parents: 40774
diff changeset
     2
    Author:     Olaf Müller
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
     3
*)
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
     4
62002
f1599e98c4d0 isabelle update_cartouches -c -t;
wenzelm
parents: 62001
diff changeset
     5
section \<open>Refinement Mappings in HOLCF/IOA\<close>
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
     6
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
     7
theory RefMappings
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
     8
imports Traces
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
     9
begin
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    10
36452
d37c6eed8117 renamed command 'defaultsort' to 'default_sort';
wenzelm
parents: 35174
diff changeset
    11
default_sort type
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    12
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    13
definition move :: "('a, 's) ioa \<Rightarrow> ('a, 's) pairs \<Rightarrow> 's \<Rightarrow> 'a \<Rightarrow> 's \<Rightarrow> bool"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    14
  where "move ioa ex s a t \<longleftrightarrow>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    15
    is_exec_frag ioa (s, ex) \<and> Finite ex \<and>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    16
    laststate (s, ex) = t \<and>
63549
b0d31c7def86 more symbols;
wenzelm
parents: 62390
diff changeset
    17
    mk_trace ioa \<cdot> ex = (if a \<in> ext ioa then a \<leadsto> nil else nil)"
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    18
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    19
definition is_ref_map :: "('s1 \<Rightarrow> 's2) \<Rightarrow> ('a, 's1) ioa \<Rightarrow> ('a, 's2) ioa \<Rightarrow> bool"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    20
  where "is_ref_map f C A \<longleftrightarrow>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    21
   ((\<forall>s \<in> starts_of C. f s \<in> starts_of A) \<and>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    22
     (\<forall>s t a. reachable C s \<and> s \<midarrow>a\<midarrow>C\<rightarrow> t \<longrightarrow> (\<exists>ex. move A ex (f s) a (f t))))"
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
    23
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    24
definition is_weak_ref_map :: "('s1 \<Rightarrow> 's2) \<Rightarrow> ('a, 's1) ioa \<Rightarrow> ('a, 's2) ioa \<Rightarrow> bool"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    25
  where "is_weak_ref_map f C A \<longleftrightarrow>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    26
    ((\<forall>s \<in> starts_of C. f s \<in> starts_of A) \<and>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    27
      (\<forall>s t a. reachable C s \<and> s \<midarrow>a\<midarrow>C\<rightarrow> t \<longrightarrow>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    28
        (if a \<in> ext C then (f s) \<midarrow>a\<midarrow>A\<rightarrow> (f t) else f s = f t)))"
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    29
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    30
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    31
subsection \<open>Transitions and moves\<close>
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    32
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    33
lemma transition_is_ex: "s \<midarrow>a\<midarrow>A\<rightarrow> t \<Longrightarrow> \<exists>ex. move A ex s a t"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    34
  apply (rule_tac x = " (a, t) \<leadsto> nil" in exI)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    35
  apply (simp add: move_def)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    36
  done
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    37
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    38
lemma nothing_is_ex: "a \<notin> ext A \<and> s = t \<Longrightarrow> \<exists>ex. move A ex s a t"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    39
  apply (rule_tac x = "nil" in exI)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    40
  apply (simp add: move_def)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    41
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    42
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    43
lemma ei_transitions_are_ex:
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    44
  "s \<midarrow>a\<midarrow>A\<rightarrow> s' \<and> s' \<midarrow>a'\<midarrow>A\<rightarrow> s'' \<and> a' \<notin> ext A \<Longrightarrow> \<exists>ex. move A ex s a s''"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    45
  apply (rule_tac x = " (a,s') \<leadsto> (a',s'') \<leadsto>nil" in exI)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    46
  apply (simp add: move_def)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    47
  done
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    48
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    49
lemma eii_transitions_are_ex:
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    50
  "s1 \<midarrow>a1\<midarrow>A\<rightarrow> s2 \<and> s2 \<midarrow>a2\<midarrow>A\<rightarrow> s3 \<and> s3 \<midarrow>a3\<midarrow>A\<rightarrow> s4 \<and> a2 \<notin> ext A \<and> a3 \<notin> ext A \<Longrightarrow>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    51
    \<exists>ex. move A ex s1 a1 s4"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    52
  apply (rule_tac x = "(a1, s2) \<leadsto> (a2, s3) \<leadsto> (a3, s4) \<leadsto> nil" in exI)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    53
  apply (simp add: move_def)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    54
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    55
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    56
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    57
subsection \<open>\<open>weak_ref_map\<close> and \<open>ref_map\<close>\<close>
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    58
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    59
lemma weak_ref_map2ref_map: "ext C = ext A \<Longrightarrow> is_weak_ref_map f C A \<Longrightarrow> is_ref_map f C A"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    60
  apply (unfold is_weak_ref_map_def is_ref_map_def)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    61
  apply auto
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    62
  apply (case_tac "a \<in> ext A")
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    63
  apply (auto intro: transition_is_ex nothing_is_ex)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    64
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    65
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    66
lemma imp_conj_lemma: "(P \<Longrightarrow> Q \<longrightarrow> R) \<Longrightarrow> P \<and> Q \<longrightarrow> R"
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    67
  by blast
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    68
62390
842917225d56 more canonical names
nipkow
parents: 62116
diff changeset
    69
declare if_split [split del]
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    70
declare if_weak_cong [cong del]
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    71
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    72
lemma rename_through_pmap:
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    73
  "is_weak_ref_map f C A \<Longrightarrow> is_weak_ref_map f (rename C g) (rename A g)"
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    74
  apply (simp add: is_weak_ref_map_def)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    75
  apply (rule conjI)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    76
  text \<open>1: start states\<close>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    77
  apply (simp add: rename_def rename_set_def starts_of_def)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    78
  text \<open>1: start states\<close>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    79
  apply (rule allI)+
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    80
  apply (rule imp_conj_lemma)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    81
  apply (simp (no_asm) add: rename_def rename_set_def)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    82
  apply (simp add: externals_def asig_inputs_def asig_outputs_def asig_of_def trans_of_def)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    83
  apply safe
62390
842917225d56 more canonical names
nipkow
parents: 62116
diff changeset
    84
  apply (simplesubst if_split)
62116
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    85
   apply (rule conjI)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    86
   apply (rule impI)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    87
   apply (erule disjE)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    88
   apply (erule exE)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    89
  apply (erule conjE)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    90
  text \<open>\<open>x\<close> is input\<close>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    91
   apply (drule sym)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    92
   apply (drule sym)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    93
  apply simp
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    94
  apply hypsubst+
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    95
  apply (frule reachable_rename)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    96
  apply simp
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    97
  text \<open>\<open>x\<close> is output\<close>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    98
   apply (erule exE)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    99
  apply (erule conjE)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   100
   apply (drule sym)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   101
   apply (drule sym)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   102
  apply simp
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   103
  apply hypsubst+
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   104
  apply (frule reachable_rename)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   105
  apply simp
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   106
  text \<open>\<open>x\<close> is internal\<close>
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   107
  apply (frule reachable_rename)
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   108
  apply auto
bc178c0fe1a1 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
   109
  done
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   110
62390
842917225d56 more canonical names
nipkow
parents: 62116
diff changeset
   111
declare if_split [split]
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
   112
declare if_weak_cong [cong]
4559
8e604d885b54 added files containing temproal logic and abstraction;
mueller
parents: 3433
diff changeset
   113
3071
981258186b71 New meta theory for IOA based on HOLCF.
mueller
parents:
diff changeset
   114
end