src/HOL/ex/Transitive_Closure_Table_Ex.thy
author wenzelm
Mon, 25 Mar 2019 17:21:26 +0100
changeset 69981 3dced198b9ec
parent 66453 cc19f7ca2ed6
permissions -rw-r--r--
more strict AFP properties;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
56922
d411a81b8356 removed junk from library theory
haftmann
parents:
diff changeset
     1
(* Author: Stefan Berghofer, Lukas Bulwahn, TU Muenchen *)
d411a81b8356 removed junk from library theory
haftmann
parents:
diff changeset
     2
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
     3
section \<open>Simple example for table-based implementation of the reflexive transitive closure\<close>
56922
d411a81b8356 removed junk from library theory
haftmann
parents:
diff changeset
     4
d411a81b8356 removed junk from library theory
haftmann
parents:
diff changeset
     5
theory Transitive_Closure_Table_Ex
66453
cc19f7ca2ed6 session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents: 66345
diff changeset
     6
imports "HOL-Library.Transitive_Closure_Table"
56922
d411a81b8356 removed junk from library theory
haftmann
parents:
diff changeset
     7
begin
d411a81b8356 removed junk from library theory
haftmann
parents:
diff changeset
     8
58310
91ea607a34d8 updated news
blanchet
parents: 58249
diff changeset
     9
datatype ty = A | B | C
56922
d411a81b8356 removed junk from library theory
haftmann
parents:
diff changeset
    10
d411a81b8356 removed junk from library theory
haftmann
parents:
diff changeset
    11
inductive test :: "ty \<Rightarrow> ty \<Rightarrow> bool"
d411a81b8356 removed junk from library theory
haftmann
parents:
diff changeset
    12
where
d411a81b8356 removed junk from library theory
haftmann
parents:
diff changeset
    13
  "test A B"
d411a81b8356 removed junk from library theory
haftmann
parents:
diff changeset
    14
| "test B A"
d411a81b8356 removed junk from library theory
haftmann
parents:
diff changeset
    15
| "test B C"
d411a81b8356 removed junk from library theory
haftmann
parents:
diff changeset
    16
d411a81b8356 removed junk from library theory
haftmann
parents:
diff changeset
    17
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    18
text \<open>Invoking with the predicate compiler and the generic code generator\<close>
56922
d411a81b8356 removed junk from library theory
haftmann
parents:
diff changeset
    19
d411a81b8356 removed junk from library theory
haftmann
parents:
diff changeset
    20
code_pred test .
d411a81b8356 removed junk from library theory
haftmann
parents:
diff changeset
    21
d411a81b8356 removed junk from library theory
haftmann
parents:
diff changeset
    22
values "{x. test\<^sup>*\<^sup>* A C}"
d411a81b8356 removed junk from library theory
haftmann
parents:
diff changeset
    23
values "{x. test\<^sup>*\<^sup>* C A}"
d411a81b8356 removed junk from library theory
haftmann
parents:
diff changeset
    24
values "{x. test\<^sup>*\<^sup>* A x}"
d411a81b8356 removed junk from library theory
haftmann
parents:
diff changeset
    25
values "{x. test\<^sup>*\<^sup>* x C}"
d411a81b8356 removed junk from library theory
haftmann
parents:
diff changeset
    26
66345
882abe912da9 do not fall back on nbe if plain evaluation fails
haftmann
parents: 61343
diff changeset
    27
value [nbe] "test\<^sup>*\<^sup>* A C"
882abe912da9 do not fall back on nbe if plain evaluation fails
haftmann
parents: 61343
diff changeset
    28
value [nbe] "test\<^sup>*\<^sup>* C A"
56922
d411a81b8356 removed junk from library theory
haftmann
parents:
diff changeset
    29
d411a81b8356 removed junk from library theory
haftmann
parents:
diff changeset
    30
end