src/HOLCF/IOA/ex/TrivEx.thy
author wenzelm
Mon, 08 Oct 2007 18:13:10 +0200
changeset 24911 4efb68e5576d
parent 19740 6b38551d0798
child 25135 4f8176c940cf
permissions -rw-r--r--
primitive add_def: strip sorts in axiom;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
6470
f3015fd68d66 moved this trivial example to new ex dir;
mueller
parents:
diff changeset
     1
(*  Title:      HOLCF/IOA/TrivEx.thy
f3015fd68d66 moved this trivial example to new ex dir;
mueller
parents:
diff changeset
     2
    ID:         $Id$
12218
wenzelm
parents: 6470
diff changeset
     3
    Author:     Olaf Müller
6470
f3015fd68d66 moved this trivial example to new ex dir;
mueller
parents:
diff changeset
     4
*)
f3015fd68d66 moved this trivial example to new ex dir;
mueller
parents:
diff changeset
     5
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
     6
header {* Trivial Abstraction Example *}
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
     7
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
     8
theory TrivEx
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
     9
imports Abstraction
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    10
begin
6470
f3015fd68d66 moved this trivial example to new ex dir;
mueller
parents:
diff changeset
    11
f3015fd68d66 moved this trivial example to new ex dir;
mueller
parents:
diff changeset
    12
datatype action = INC
f3015fd68d66 moved this trivial example to new ex dir;
mueller
parents:
diff changeset
    13
f3015fd68d66 moved this trivial example to new ex dir;
mueller
parents:
diff changeset
    14
consts
f3015fd68d66 moved this trivial example to new ex dir;
mueller
parents:
diff changeset
    15
f3015fd68d66 moved this trivial example to new ex dir;
mueller
parents:
diff changeset
    16
C_asig   ::  "action signature"
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    17
C_trans  :: "(action, nat)transition set"
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    18
C_ioa    :: "(action, nat)ioa"
6470
f3015fd68d66 moved this trivial example to new ex dir;
mueller
parents:
diff changeset
    19
f3015fd68d66 moved this trivial example to new ex dir;
mueller
parents:
diff changeset
    20
A_asig   :: "action signature"
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    21
A_trans  :: "(action, bool)transition set"
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    22
A_ioa    :: "(action, bool)ioa"
6470
f3015fd68d66 moved this trivial example to new ex dir;
mueller
parents:
diff changeset
    23
f3015fd68d66 moved this trivial example to new ex dir;
mueller
parents:
diff changeset
    24
h_abs    :: "nat => bool"
f3015fd68d66 moved this trivial example to new ex dir;
mueller
parents:
diff changeset
    25
f3015fd68d66 moved this trivial example to new ex dir;
mueller
parents:
diff changeset
    26
defs
f3015fd68d66 moved this trivial example to new ex dir;
mueller
parents:
diff changeset
    27
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    28
C_asig_def:
6470
f3015fd68d66 moved this trivial example to new ex dir;
mueller
parents:
diff changeset
    29
  "C_asig == ({},{INC},{})"
f3015fd68d66 moved this trivial example to new ex dir;
mueller
parents:
diff changeset
    30
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    31
C_trans_def: "C_trans ==
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    32
 {tr. let s = fst(tr);
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    33
          t = snd(snd(tr))
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    34
      in case fst(snd(tr))
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    35
      of
6470
f3015fd68d66 moved this trivial example to new ex dir;
mueller
parents:
diff changeset
    36
      INC       => t = Suc(s)}"
f3015fd68d66 moved this trivial example to new ex dir;
mueller
parents:
diff changeset
    37
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    38
C_ioa_def: "C_ioa ==
6470
f3015fd68d66 moved this trivial example to new ex dir;
mueller
parents:
diff changeset
    39
 (C_asig, {0}, C_trans,{},{})"
f3015fd68d66 moved this trivial example to new ex dir;
mueller
parents:
diff changeset
    40
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    41
A_asig_def:
6470
f3015fd68d66 moved this trivial example to new ex dir;
mueller
parents:
diff changeset
    42
  "A_asig == ({},{INC},{})"
f3015fd68d66 moved this trivial example to new ex dir;
mueller
parents:
diff changeset
    43
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    44
A_trans_def: "A_trans ==
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    45
 {tr. let s = fst(tr);
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    46
          t = snd(snd(tr))
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    47
      in case fst(snd(tr))
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    48
      of
6470
f3015fd68d66 moved this trivial example to new ex dir;
mueller
parents:
diff changeset
    49
      INC       => t = True}"
f3015fd68d66 moved this trivial example to new ex dir;
mueller
parents:
diff changeset
    50
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    51
A_ioa_def: "A_ioa ==
6470
f3015fd68d66 moved this trivial example to new ex dir;
mueller
parents:
diff changeset
    52
 (A_asig, {False}, A_trans,{},{})"
f3015fd68d66 moved this trivial example to new ex dir;
mueller
parents:
diff changeset
    53
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    54
h_abs_def:
6470
f3015fd68d66 moved this trivial example to new ex dir;
mueller
parents:
diff changeset
    55
  "h_abs n == n~=0"
f3015fd68d66 moved this trivial example to new ex dir;
mueller
parents:
diff changeset
    56
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    57
axioms
6470
f3015fd68d66 moved this trivial example to new ex dir;
mueller
parents:
diff changeset
    58
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    59
MC_result:
6470
f3015fd68d66 moved this trivial example to new ex dir;
mueller
parents:
diff changeset
    60
  "validIOA A_ioa (<>[] <%(b,a,c). b>)"
f3015fd68d66 moved this trivial example to new ex dir;
mueller
parents:
diff changeset
    61
19740
6b38551d0798 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    62
lemma h_abs_is_abstraction:
6b38551d0798 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    63
  "is_abstraction h_abs C_ioa A_ioa"
6b38551d0798 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    64
apply (unfold is_abstraction_def)
6b38551d0798 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    65
apply (rule conjI)
6b38551d0798 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    66
txt {* start states *}
6b38551d0798 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    67
apply (simp (no_asm) add: h_abs_def starts_of_def C_ioa_def A_ioa_def)
6b38551d0798 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    68
txt {* step case *}
6b38551d0798 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    69
apply (rule allI)+
6b38551d0798 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    70
apply (rule imp_conj_lemma)
6b38551d0798 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    71
apply (simp (no_asm) add: trans_of_def C_ioa_def A_ioa_def C_trans_def A_trans_def)
6b38551d0798 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    72
apply (induct_tac "a")
6b38551d0798 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    73
apply (simp add: h_abs_def)
6b38551d0798 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    74
done
6b38551d0798 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    75
6b38551d0798 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    76
lemma TrivEx_abstraction: "validIOA C_ioa (<>[] <%(n,a,m). n~=0>)"
6b38551d0798 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    77
apply (rule AbsRuleT1)
6b38551d0798 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    78
apply (rule h_abs_is_abstraction)
6b38551d0798 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    79
apply (rule MC_result)
6b38551d0798 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    80
apply (tactic "abstraction_tac 1")
6b38551d0798 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    81
apply (simp add: h_abs_def)
6b38551d0798 removed legacy ML scripts;
wenzelm
parents: 17244
diff changeset
    82
done
17244
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    83
0b2ff9541727 converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    84
end