src/HOL/HOLCF/IOA/Simulations.thy
author haftmann
Thu, 19 Jun 2025 17:15:40 +0200
changeset 82734 89347c0cc6a3
parent 67613 ce654b0e6d69
permissions -rw-r--r--
treat map_filter similar to list_all, list_ex, list_ex1
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/Simulations.thy
40945
b8703f63bfb2 recoded latin1 as utf8;
wenzelm
parents: 40774
diff changeset
     2
    Author:     Olaf Müller
4565
ea467ce15040 added forward simulation correectness;
mueller
parents:
diff changeset
     3
*)
ea467ce15040 added forward simulation correectness;
mueller
parents:
diff changeset
     4
62002
f1599e98c4d0 isabelle update_cartouches -c -t;
wenzelm
parents: 58880
diff changeset
     5
section \<open>Simulations in HOLCF/IOA\<close>
4565
ea467ce15040 added forward simulation correectness;
mueller
parents:
diff changeset
     6
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
     7
theory Simulations
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
     8
imports RefCorrectness
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
4565
ea467ce15040 added forward simulation correectness;
mueller
parents:
diff changeset
    12
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    13
definition is_simulation :: "('s1 \<times> 's2) set \<Rightarrow> ('a, 's1) ioa \<Rightarrow> ('a, 's2) ioa \<Rightarrow> bool"
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    14
  where "is_simulation R C A \<longleftrightarrow>
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    15
    (\<forall>s \<in> starts_of C. R``{s} \<inter> starts_of A \<noteq> {}) \<and>
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    16
    (\<forall>s s' t a. reachable C s \<and> s \<midarrow>a\<midarrow>C\<rightarrow> t \<and> (s, s') \<in> R
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    17
      \<longrightarrow> (\<exists>t' ex. (t, t') \<in> R \<and> move A ex s' a t'))"
4565
ea467ce15040 added forward simulation correectness;
mueller
parents:
diff changeset
    18
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    19
definition is_backward_simulation :: "('s1 \<times> 's2) set \<Rightarrow> ('a, 's1) ioa \<Rightarrow> ('a, 's2) ioa \<Rightarrow> bool"
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    20
  where "is_backward_simulation R C A \<longleftrightarrow>
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    21
    (\<forall>s \<in> starts_of C. R``{s} \<subseteq> starts_of A) \<and>
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    22
    (\<forall>s t t' a. reachable C s \<and> s \<midarrow>a\<midarrow>C\<rightarrow> t \<and> (t, t') \<in> R
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    23
      \<longrightarrow> (\<exists>ex s'. (s,s') \<in> R \<and> move A ex s' a t'))"
4565
ea467ce15040 added forward simulation correectness;
mueller
parents:
diff changeset
    24
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    25
definition is_forw_back_simulation ::
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    26
    "('s1 \<times> 's2 set) set \<Rightarrow> ('a, 's1) ioa \<Rightarrow> ('a, 's2) ioa \<Rightarrow> bool"
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    27
  where "is_forw_back_simulation R C A \<longleftrightarrow>
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    28
    (\<forall>s \<in> starts_of C. \<exists>S'. (s, S') \<in> R \<and> S' \<subseteq> starts_of A) \<and>
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    29
    (\<forall>s S' t a. reachable C s \<and> s \<midarrow>a\<midarrow>C\<rightarrow> t \<and> (s, S') \<in> R
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    30
      \<longrightarrow> (\<exists>T'. (t, T') \<in> R \<and> (\<forall>t' \<in> T'. \<exists>s' \<in> S'. \<exists>ex. move A ex s' a t')))"
4565
ea467ce15040 added forward simulation correectness;
mueller
parents:
diff changeset
    31
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    32
definition is_back_forw_simulation ::
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    33
    "('s1 \<times> 's2 set) set \<Rightarrow> ('a, 's1) ioa \<Rightarrow> ('a, 's2) ioa \<Rightarrow> bool"
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    34
  where "is_back_forw_simulation R C A \<longleftrightarrow>
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    35
    ((\<forall>s \<in> starts_of C. \<forall>S'. (s, S') \<in> R \<longrightarrow> S' \<inter> starts_of A \<noteq> {}) \<and>
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    36
    (\<forall>s t T' a. reachable C s \<and> s \<midarrow>a\<midarrow>C\<rightarrow> t \<and> (t, T') \<in> R
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    37
      \<longrightarrow> (\<exists>S'. (s, S') \<in> R \<and> (\<forall>s' \<in> S'. \<exists>t' \<in> T'. \<exists>ex. move A ex s' a t'))))"
4565
ea467ce15040 added forward simulation correectness;
mueller
parents:
diff changeset
    38
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    39
definition is_history_relation :: "('s1 \<times> 's2) set \<Rightarrow> ('a, 's1) ioa \<Rightarrow> ('a, 's2) ioa \<Rightarrow> bool"
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    40
  where "is_history_relation R C A \<longleftrightarrow>
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62192
diff changeset
    41
    is_simulation R C A \<and> is_ref_map (\<lambda>x. (SOME y. (x, y) \<in> R\<inverse>)) A C"
4565
ea467ce15040 added forward simulation correectness;
mueller
parents:
diff changeset
    42
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    43
definition is_prophecy_relation :: "('s1 \<times> 's2) set \<Rightarrow> ('a, 's1) ioa \<Rightarrow> ('a, 's2) ioa \<Rightarrow> bool"
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    44
  where "is_prophecy_relation R C A \<longleftrightarrow>
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62192
diff changeset
    45
    is_backward_simulation R C A \<and> is_ref_map (\<lambda>x. (SOME y. (x, y) \<in> R\<inverse>)) A C"
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    46
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    47
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    48
lemma set_non_empty: "A \<noteq> {} \<longleftrightarrow> (\<exists>x. x \<in> A)"
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    49
  by auto
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    50
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    51
lemma Int_non_empty: "A \<inter> B \<noteq> {} \<longleftrightarrow> (\<exists>x. x \<in> A \<and> x \<in> B)"
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    52
  by (simp add: set_non_empty)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    53
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    54
lemma Sim_start_convert [simp]: "R``{x} \<inter> S \<noteq> {} \<longleftrightarrow> (\<exists>y. (x, y) \<in> R \<and> y \<in> S)"
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    55
  by (simp add: Image_def Int_non_empty)
19741
f65265d71426 removed legacy ML scripts;
wenzelm
parents: 17233
diff changeset
    56
62192
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    57
lemma ref_map_is_simulation: "is_ref_map f C A \<Longrightarrow> is_simulation {p. snd p = f (fst p)} C A"
bdaa0eb0fc74 misc tuning and modernization;
wenzelm
parents: 62008
diff changeset
    58
  by (simp add: is_ref_map_def is_simulation_def)
17233
41eee2e7b465 converted specifications to Isar theories;
wenzelm
parents: 14981
diff changeset
    59
4565
ea467ce15040 added forward simulation correectness;
mueller
parents:
diff changeset
    60
end