src/HOL/Prolog/Test.thy
author wenzelm
Wed, 28 Aug 2024 22:54:45 +0200
changeset 80786 70076ba563d2
parent 69597 ff784d5a5bfb
child 80914 d97fdabd9e2b
permissions -rw-r--r--
more specific "args" syntax, to support more markup for syntax consts;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13208
965f95a3abd9 added the usual file headers
oheimb
parents: 12338
diff changeset
     1
(*  Title:    HOL/Prolog/Test.thy
965f95a3abd9 added the usual file headers
oheimb
parents: 12338
diff changeset
     2
    Author:   David von Oheimb (based on a lecture on Lambda Prolog by Nadathur)
965f95a3abd9 added the usual file headers
oheimb
parents: 12338
diff changeset
     3
*)
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
     4
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61337
diff changeset
     5
section \<open>Basic examples\<close>
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
     6
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
     7
theory Test
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
     8
imports HOHH
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
     9
begin
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    10
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    11
typedecl nat
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    12
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    13
typedecl 'a list
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    14
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    15
consts
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    16
  Nil   :: "'a list"                                  ("[]")
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    17
  Cons  :: "'a => 'a list => 'a list"                 (infixr "#"  65)
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    18
80786
70076ba563d2 more specific "args" syntax, to support more markup for syntax consts;
wenzelm
parents: 69597
diff changeset
    19
text \<open>List enumeration\<close>
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    20
80786
70076ba563d2 more specific "args" syntax, to support more markup for syntax consts;
wenzelm
parents: 69597
diff changeset
    21
nonterminal list_args
70076ba563d2 more specific "args" syntax, to support more markup for syntax consts;
wenzelm
parents: 69597
diff changeset
    22
syntax
70076ba563d2 more specific "args" syntax, to support more markup for syntax consts;
wenzelm
parents: 69597
diff changeset
    23
  "" :: "'a \<Rightarrow> list_args"  ("_")
70076ba563d2 more specific "args" syntax, to support more markup for syntax consts;
wenzelm
parents: 69597
diff changeset
    24
  "_list_args" :: "'a \<Rightarrow> list_args \<Rightarrow> list_args"  ("_,/ _")
70076ba563d2 more specific "args" syntax, to support more markup for syntax consts;
wenzelm
parents: 69597
diff changeset
    25
  "_list" :: "list_args => 'a list"    ("[(_)]")
70076ba563d2 more specific "args" syntax, to support more markup for syntax consts;
wenzelm
parents: 69597
diff changeset
    26
syntax_consts
70076ba563d2 more specific "args" syntax, to support more markup for syntax consts;
wenzelm
parents: 69597
diff changeset
    27
  "_list_args" "_list" == Cons
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    28
translations
80786
70076ba563d2 more specific "args" syntax, to support more markup for syntax consts;
wenzelm
parents: 69597
diff changeset
    29
  "[x, xs]" == "x#[xs]"
70076ba563d2 more specific "args" syntax, to support more markup for syntax consts;
wenzelm
parents: 69597
diff changeset
    30
  "[x]" == "x#[]"
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    31
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    32
typedecl person
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    33
51311
337cfc42c9c8 eliminated legacy 'axioms';
wenzelm
parents: 46473
diff changeset
    34
axiomatization
337cfc42c9c8 eliminated legacy 'axioms';
wenzelm
parents: 46473
diff changeset
    35
  append  :: "['a list, 'a list, 'a list]            => bool" and
337cfc42c9c8 eliminated legacy 'axioms';
wenzelm
parents: 46473
diff changeset
    36
  reverse :: "['a list, 'a list]                     => bool" and
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    37
51311
337cfc42c9c8 eliminated legacy 'axioms';
wenzelm
parents: 46473
diff changeset
    38
  mappred :: "[('a => 'b => bool), 'a list, 'b list] => bool" and
337cfc42c9c8 eliminated legacy 'axioms';
wenzelm
parents: 46473
diff changeset
    39
  mapfun  :: "[('a => 'b), 'a list, 'b list]         => bool" and
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    40
51311
337cfc42c9c8 eliminated legacy 'axioms';
wenzelm
parents: 46473
diff changeset
    41
  bob     :: person and
337cfc42c9c8 eliminated legacy 'axioms';
wenzelm
parents: 46473
diff changeset
    42
  sue     :: person and
337cfc42c9c8 eliminated legacy 'axioms';
wenzelm
parents: 46473
diff changeset
    43
  ned     :: person and
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    44
51311
337cfc42c9c8 eliminated legacy 'axioms';
wenzelm
parents: 46473
diff changeset
    45
  nat23   :: nat          ("23") and
337cfc42c9c8 eliminated legacy 'axioms';
wenzelm
parents: 46473
diff changeset
    46
  nat24   :: nat          ("24") and
337cfc42c9c8 eliminated legacy 'axioms';
wenzelm
parents: 46473
diff changeset
    47
  nat25   :: nat          ("25") and
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    48
51311
337cfc42c9c8 eliminated legacy 'axioms';
wenzelm
parents: 46473
diff changeset
    49
  age     :: "[person, nat]                          => bool" and
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    50
51311
337cfc42c9c8 eliminated legacy 'axioms';
wenzelm
parents: 46473
diff changeset
    51
  eq      :: "['a, 'a]                               => bool" and
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    52
51311
337cfc42c9c8 eliminated legacy 'axioms';
wenzelm
parents: 46473
diff changeset
    53
  empty   :: "['b]                                   => bool" and
337cfc42c9c8 eliminated legacy 'axioms';
wenzelm
parents: 46473
diff changeset
    54
  add     :: "['a, 'b, 'b]                           => bool" and
337cfc42c9c8 eliminated legacy 'axioms';
wenzelm
parents: 46473
diff changeset
    55
  remove  :: "['a, 'b, 'b]                           => bool" and
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    56
  bag_appl:: "['a, 'a, 'a, 'a]                       => bool"
51311
337cfc42c9c8 eliminated legacy 'axioms';
wenzelm
parents: 46473
diff changeset
    57
where
337cfc42c9c8 eliminated legacy 'axioms';
wenzelm
parents: 46473
diff changeset
    58
        append:  "\<And>x xs ys zs. append  []    xs  xs    ..
337cfc42c9c8 eliminated legacy 'axioms';
wenzelm
parents: 46473
diff changeset
    59
                  append (x#xs) ys (x#zs) :- append xs ys zs" and
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
    60
        reverse: "\<And>L1 L2. reverse L1 L2 :- (\<forall>rev_aux.
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
    61
                  (\<forall>L.          rev_aux  []    L  L )..
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
    62
                  (\<forall>X L1 L2 L3. rev_aux (X#L1) L2 L3 :- rev_aux L1 L2 (X#L3))
51311
337cfc42c9c8 eliminated legacy 'axioms';
wenzelm
parents: 46473
diff changeset
    63
                  => rev_aux L1 L2 [])" and
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    64
51311
337cfc42c9c8 eliminated legacy 'axioms';
wenzelm
parents: 46473
diff changeset
    65
        mappred: "\<And>x xs y ys P. mappred P  []     []    ..
337cfc42c9c8 eliminated legacy 'axioms';
wenzelm
parents: 46473
diff changeset
    66
                  mappred P (x#xs) (y#ys) :- P x y  &  mappred P xs ys" and
337cfc42c9c8 eliminated legacy 'axioms';
wenzelm
parents: 46473
diff changeset
    67
        mapfun:  "\<And>x xs ys f. mapfun f  []     []      ..
337cfc42c9c8 eliminated legacy 'axioms';
wenzelm
parents: 46473
diff changeset
    68
                  mapfun f (x#xs) (f x#ys) :- mapfun f xs ys" and
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    69
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    70
        age:     "age bob 24 ..
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    71
                  age sue 23 ..
51311
337cfc42c9c8 eliminated legacy 'axioms';
wenzelm
parents: 46473
diff changeset
    72
                  age ned 23" and
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    73
51311
337cfc42c9c8 eliminated legacy 'axioms';
wenzelm
parents: 46473
diff changeset
    74
        eq:      "\<And>x. eq x x" and
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    75
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    76
(* actual definitions of empty and add is hidden -> yields abstract data type *)
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    77
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
    78
        bag_appl: "\<And>A B X Y. bag_appl A B X Y:- (\<exists>S1 S2 S3 S4 S5.
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    79
                                empty    S1    &
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    80
                                add A    S1 S2 &
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    81
                                add B    S2 S3 &
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    82
                                remove X S3 S4 &
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    83
                                remove Y S4 S5 &
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    84
                                empty    S5)"
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    85
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
    86
lemmas prog_Test = append reverse mappred mapfun age eq bag_appl
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
    87
61337
4645502c3c64 fewer aliases for toplevel theorem statements;
wenzelm
parents: 58889
diff changeset
    88
schematic_goal "append ?x ?y [a,b,c,d]"
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
    89
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
    90
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
    91
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
    92
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
    93
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
    94
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
    95
61337
4645502c3c64 fewer aliases for toplevel theorem statements;
wenzelm
parents: 58889
diff changeset
    96
schematic_goal "append [a,b] y ?L"
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
    97
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
    98
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
    99
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   100
schematic_goal "\<forall>y. append [a,b] y (?L y)"
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   101
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   102
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   103
61337
4645502c3c64 fewer aliases for toplevel theorem statements;
wenzelm
parents: 58889
diff changeset
   104
schematic_goal "reverse [] ?L"
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   105
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   106
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   107
61337
4645502c3c64 fewer aliases for toplevel theorem statements;
wenzelm
parents: 58889
diff changeset
   108
schematic_goal "reverse [23] ?L"
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   109
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   110
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   111
61337
4645502c3c64 fewer aliases for toplevel theorem statements;
wenzelm
parents: 58889
diff changeset
   112
schematic_goal "reverse [23,24,?x] ?L"
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   113
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   114
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   115
61337
4645502c3c64 fewer aliases for toplevel theorem statements;
wenzelm
parents: 58889
diff changeset
   116
schematic_goal "reverse ?L [23,24,?x]"
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   117
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   118
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   119
61337
4645502c3c64 fewer aliases for toplevel theorem statements;
wenzelm
parents: 58889
diff changeset
   120
schematic_goal "mappred age ?x [23,24]"
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   121
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   122
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   123
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   124
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   125
schematic_goal "mappred (\<lambda>x y. \<exists>z. age z y) ?x [23,24]"
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   126
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   127
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   128
61337
4645502c3c64 fewer aliases for toplevel theorem statements;
wenzelm
parents: 58889
diff changeset
   129
schematic_goal "mappred ?P [bob,sue] [24,23]"
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   130
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   131
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   132
61337
4645502c3c64 fewer aliases for toplevel theorem statements;
wenzelm
parents: 58889
diff changeset
   133
schematic_goal "mapfun f [bob,bob,sue] [?x,?y,?z]"
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   134
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   135
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   136
61337
4645502c3c64 fewer aliases for toplevel theorem statements;
wenzelm
parents: 58889
diff changeset
   137
schematic_goal "mapfun (%x. h x 25) [bob,sue] ?L"
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   138
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   139
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   140
61337
4645502c3c64 fewer aliases for toplevel theorem statements;
wenzelm
parents: 58889
diff changeset
   141
schematic_goal "mapfun ?F [24,25] [h bob 24,h bob 25]"
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   142
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   143
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   144
61337
4645502c3c64 fewer aliases for toplevel theorem statements;
wenzelm
parents: 58889
diff changeset
   145
schematic_goal "mapfun ?F [24] [h 24 24]"
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   146
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   147
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   148
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   149
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   150
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   151
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   152
lemma "True"
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   153
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   154
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   155
61337
4645502c3c64 fewer aliases for toplevel theorem statements;
wenzelm
parents: 58889
diff changeset
   156
schematic_goal "age ?x 24 & age ?y 23"
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   157
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   158
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   159
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   160
61337
4645502c3c64 fewer aliases for toplevel theorem statements;
wenzelm
parents: 58889
diff changeset
   161
schematic_goal "age ?x 24 | age ?x 23"
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   162
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   163
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   164
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   165
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   166
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   167
lemma "\<exists>x y. age x y"
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   168
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   169
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   170
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   171
lemma "\<forall>x y. append [] x x"
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   172
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   173
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   174
61337
4645502c3c64 fewer aliases for toplevel theorem statements;
wenzelm
parents: 58889
diff changeset
   175
schematic_goal "age sue 24 .. age bob 23 => age ?x ?y"
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   176
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   177
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   178
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   179
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   180
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   181
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   182
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   183
(*set trace_DEPTH_FIRST;*)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   184
lemma "age bob 25 :- age bob 24 => age bob 25"
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   185
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   186
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   187
(*reset trace_DEPTH_FIRST;*)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   188
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   189
schematic_goal "(\<forall>x. age x 25 :- age x 23) => age ?x 25 & age ?y 25"
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   190
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   191
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   192
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   193
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   194
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   195
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   196
lemma "\<forall>x. \<exists>y. eq x y"
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   197
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   198
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   199
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   200
schematic_goal "\<exists>P. P \<and> eq P ?x"
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   201
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   202
(*
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   203
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   204
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   205
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   206
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   207
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   208
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   209
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   210
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   211
*)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   212
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   213
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   214
lemma "\<exists>P. eq P (True & True) \<and> P"
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   215
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   216
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   217
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   218
lemma "\<exists>P. eq P (\<or>) \<and> P k True"
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   219
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   220
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   221
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   222
lemma "\<exists>P. eq P (Q => True) \<and> P"
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   223
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   224
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   225
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   226
(* P flexible: *)
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   227
lemma "(\<forall>P k l. P k l :- eq P Q) => Q a b"
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   228
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   229
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   230
(*
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   231
loops:
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   232
lemma "(!P k l. P k l :- eq P (%x y. x | y)) => a | b"
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   233
*)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   234
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   235
(* implication and disjunction in atom: *)
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   236
lemma "\<exists>Q. (\<forall>p q. R(q :- p) => R(Q p q)) \<and> Q (t | s) (s | t)"
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   237
  by fast
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   238
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   239
(* disjunction in atom: *)
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   240
lemma "(\<forall>P. g P :- (P => b \<or> a)) => g(a \<or> b)"
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   241
  apply (tactic "step_tac (put_claset HOL_cs \<^context>) 1")
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   242
  apply (tactic "step_tac (put_claset HOL_cs \<^context>) 1")
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   243
  apply (tactic "step_tac (put_claset HOL_cs \<^context>) 1")
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   244
   prefer 2
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   245
   apply fast
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   246
  apply fast
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   247
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   248
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   249
(*
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   250
hangs:
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   251
lemma "(!P. g P :- (P => b | a)) => g(a | b)"
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   252
  by fast
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   253
*)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   254
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   255
schematic_goal "\<forall>Emp Stk.(
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   256
                       empty    (Emp::'b) .. 
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   257
         (\<forall>(X::nat) S. add    X (S::'b)         (Stk X S)) .. 
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   258
         (\<forall>(X::nat) S. remove X ((Stk X S)::'b) S))
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   259
 => bag_appl 23 24 ?X ?Y"
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   260
  oops
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   261
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   262
schematic_goal "\<forall>Qu. ( 
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   263
          (\<forall>L.            empty    (Qu L L)) .. 
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   264
          (\<forall>(X::nat) L K. add    X (Qu L (X#K)) (Qu L K)) ..
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   265
          (\<forall>(X::nat) L K. remove X (Qu (X#L) K) (Qu L K)))
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   266
 => bag_appl 23 24 ?X ?Y"
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   267
  oops
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   268
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   269
lemma "D \<and> (\<forall>y. E) :- (\<forall>x. True \<and> True) :- True => D"
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   270
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   271
  done
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   272
61337
4645502c3c64 fewer aliases for toplevel theorem statements;
wenzelm
parents: 58889
diff changeset
   273
schematic_goal "P x .. P y => P ?X"
21425
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   274
  apply (prolog prog_Test)
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   275
  back
c11ab38b78a7 HOL-Prolog: converted legacy ML scripts;
wenzelm
parents: 17311
diff changeset
   276
  done
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
   277
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
   278
end