src/HOL/IMP/Star.thy
author nipkow
Wed, 14 Sep 2011 06:49:01 +0200
changeset 44923 b80108b346a9
parent 43141 11fce8564415
child 45015 fdac1e9880eb
permissions -rw-r--r--
cleand up AbsInt fixpoint iteration; tuned syntax
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
43141
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
     1
theory Star imports Main
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
     2
begin
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
     3
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
     4
inductive
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
     5
  star :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
     6
for r where
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
     7
refl:  "star r x x" |
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
     8
step:  "r x y \<Longrightarrow> star r y z \<Longrightarrow> star r x z"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
     9
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    10
lemma star_trans:
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    11
  "star r x y \<Longrightarrow> star r y z \<Longrightarrow> star r x z"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    12
proof(induct rule: star.induct)
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    13
  case refl thus ?case .
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    14
next
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    15
  case step thus ?case by (metis star.step)
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    16
qed
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    17
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    18
lemmas star_induct = star.induct[of "r:: 'a*'b \<Rightarrow> 'a*'b \<Rightarrow> bool", split_format(complete)]
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    19
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    20
declare star.refl[simp,intro]
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    21
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    22
lemma step1[simp, intro]: "r x y \<Longrightarrow> star r x y"
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    23
by(metis refl step)
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    24
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    25
code_pred star .
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    26
11fce8564415 Replacing old IMP with new Semantics material
nipkow
parents:
diff changeset
    27
end