src/HOL/MicroJava/BV/Typing_Framework.thy
author oheimb
Tue, 12 Jun 2001 14:11:00 +0200
changeset 11372 648795477bb5
parent 11299 1d3d110c456e
child 12516 d09d0f160888
permissions -rw-r--r--
corrected xsymbol/HTML syntax
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
11229
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
     1
(*  Title:      HOL/MicroJava/BV/Typing_Framework.thy
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
     2
    ID:         $Id$
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
     3
    Author:     Tobias Nipkow
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
     4
    Copyright   2000 TUM
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
     5
*)
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
     6
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
     7
header "Typing and Dataflow Analysis Framework"
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
     8
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
     9
theory Typing_Framework = Listn:
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    10
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    11
text {* 
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    12
  The relationship between dataflow analysis and a welltyped-insruction predicate. 
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    13
*}
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    14
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    15
constdefs
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    16
 bounded :: "(nat => nat list) => nat => bool"
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    17
"bounded succs n == !p<n. !q:set(succs p). q<n"
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    18
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    19
 stable :: "'s ord =>
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    20
           (nat => 's => 's)
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    21
           => (nat => nat list) => 's list => nat => bool"
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    22
"stable r step succs ss p == !q:set(succs p). step p (ss!p) <=_r ss!q"
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    23
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    24
 stables :: "'s ord => (nat => 's => 's)
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    25
            => (nat => nat list) => 's list => bool"
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    26
"stables r step succs ss == !p<size ss. stable r step succs ss p"
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    27
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    28
 is_bcv :: "'s ord => 's => (nat => 's => 's) => (nat => nat list)
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    29
           => nat => 's set => ('s list => 's list) => bool"
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    30
"is_bcv r T step succs n A bcv == !ss : list n A.
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    31
   (!p<n. (bcv ss)!p ~= T) =
11299
1d3d110c456e welltyping -> wt_step
nipkow
parents: 11229
diff changeset
    32
   (? ts: list n A. ss <=[r] ts & wt_step r T step succs ts)"
11229
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    33
11299
1d3d110c456e welltyping -> wt_step
nipkow
parents: 11229
diff changeset
    34
 wt_step ::
11229
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    35
"'s ord => 's => (nat => 's => 's) => (nat => nat list) => 's list => bool"
11299
1d3d110c456e welltyping -> wt_step
nipkow
parents: 11229
diff changeset
    36
"wt_step r T step succs ts ==
11229
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    37
 !p<size(ts). ts!p ~= T & stable r step succs ts p"
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    38
f417841385b7 Got rid of is_dfa
nipkow
parents:
diff changeset
    39
end